Library expr_b_dp

Load seplog_header.

Require Import Bool.
Fixpoint inb (A:Set) (B:A->A->bool) (l:list A) (a:A) {struct l} : bool :=
  match l with
    nil => false
    | hd :: tl => if B a hd then true else inb A B tl a
  end.
Implicit Arguments inb.

Fixpoint add_elt_list (A:Set) (B:A->A->bool) (l:list A) (a:A) {struct l} : list A :=
  match l with
    nil => a :: nil
    | hd :: tl => if B hd a then l else hd :: add_elt_list A B tl a
  end.
Implicit Arguments add_elt_list.

Fixpoint app_list (A:Set) (B:A->A->bool) (l1 l2:list A) {struct l1} : list A :=
  match l1 with
    nil => l2
    | hd :: tl => app_list A B tl (add_elt_list B l2 hd)
  end.
Implicit Arguments app_list.

Definition option_app (A:Set) (l1 l2:option (list A)) : option (list A) :=
  match l1 with
    None => None
    | Some l1' => match l2 with
                    None => None
                    | Some l2' => Some (l1' ++ l2')
                  end
  end.
Implicit Arguments option_app.

Definition andb_option (a b: option bool) :=
  match a with
    None => None
    | Some a' => match b with
                   None => None
                   | Some b' => Some (andb a' b')
                 end
  end.

Definition orb_option (a b: option bool) :=
  match a with
    None => None
    | Some a' => match b with
                   None => None
                   | Some b' => Some (orb a' b')
                 end
  end.

Notation "e << e'" := (neg_b (ge_b e e')) (at level 78) : sep_scope.
Notation "e <<= e'" := (neg_b (gt_b e e')) (at level 78) : sep_scope.
Notation "! e" := (neg_b e) (at level 78) : sep_scope.
Notation " s '|b=' b " := (eval_b b s = true) (right associativity, at level 80).
Notation " s '|b!=' b " := (eval_b b s = false) (right associativity, at level 80).
Notation " b1 =b> b2 " := ((!b1) ||| b2) (right associativity, at level 80).

Fixpoint Expr_var (e:expr) {struct e} : list var.v :=
  match e with
    var_e x => x::nil
    | int_e z => nil
    | add_e e1 e2 => app_list beq_nat (Expr_var e1) (Expr_var e2)
    | min_e e1 e2 => app_list beq_nat (Expr_var e1) (Expr_var e2)
    | mul_e e1 e2 => app_list beq_nat (Expr_var e1) (Expr_var e2)
    | div_e e1 e2 => app_list beq_nat (Expr_var e1) (Expr_var e2)
    | mod_e e1 e2 => app_list beq_nat (Expr_var e1) (Expr_var e2)
  end.

Fixpoint Expr_B_var (e:expr_b) : list var.v :=
  match e with
    true_b => nil
    | eq_b e1 e2 => app_list beq_nat (Expr_var e1) (Expr_var e2)
    | neq_b e1 e2 => app_list beq_nat (Expr_var e1) (Expr_var e2)
    | ge_b e1 e2 => app_list beq_nat (Expr_var e1) (Expr_var e2)
    | gt_b e1 e2 => app_list beq_nat (Expr_var e1) (Expr_var e2)
    | and_b e1 e2 => app_list beq_nat (Expr_B_var e1) (Expr_B_var e2)
    | or_b e1 e2 => app_list beq_nat (Expr_B_var e1) (Expr_B_var e2)
    | neg_b e => Expr_B_var e
  end.

Lemma expr_b_neg_involutive: forall a,
 (forall s, eval_b (!! a) s = eval_b a s).

Lemma eval_b_destruct: forall b s,
 eval_b b s = true \/ eval_b b s = false.

Lemma expr_b_eq: forall b1 b2 s,
  ((s |b= b1) <-> (s |b= b2)) -> (eval_b b1 s = eval_b b2 s).

Lemma expr_b_eq': forall b1 b2 s,
  (eval_b b1 s = eval_b b2 s) -> ((s |b= b1) <-> (s |b= b2)).

Fixpoint Expr_B_size (e:expr_b) : nat :=
  match e with
    true_b => 1
    | eq_b e1 e2 => 1
    | neq_b e1 e2 => 1
    | ge_b e1 e2 => 1
    | gt_b e1 e2 => 1
    | and_b e1 e2 => 1 + (Expr_B_size e1) + (Expr_B_size e2)
    | or_b e1 e2 => 1 + (Expr_B_size e1) + (Expr_B_size e2)
    | neg_b e => 1 + Expr_B_size e
  end.

Lemma expr_b_min_size: forall b,
  Expr_B_size b >= 1.


Lemma neg_propagate_preserve: forall b n,
  (forall s, eval_b (neg_propagate b n) s = eval_b (if n then (neg_b b) else b) s).

Inductive is_neg_propagate : expr_b -> Prop :=

| true_b_is_neg_propagate: is_neg_propagate true_b
| eq_b_is_neg_propagate: forall e1 e2, is_neg_propagate (eq_b e1 e2)
| neq_b_is_neg_propagate: forall e1 e2, is_neg_propagate (neq_b e1 e2)
| ge_b_is_neg_propagate: forall e1 e2, is_neg_propagate (ge_b e1 e2)
| gt_b_is_neg_propagate: forall e1 e2, is_neg_propagate (gt_b e1 e2)

| neg_b_is_neg_propagate: forall e, Expr_B_size e = 1 -> is_neg_propagate (neg_b e)

| and_b_is_neg_propagate: forall e1 e2,
  (is_neg_propagate e1) ->
  (is_neg_propagate e2) ->
  (is_neg_propagate (and_b e1 e2))

| or_is_neg_propagate: forall e1 e2,
  (is_neg_propagate e1) ->
  (is_neg_propagate e2) ->
  (is_neg_propagate (or_b e1 e2)).

Lemma neg_propagate_correct: forall b n,
  is_neg_propagate (neg_propagate b n).

Definition constraint := expr.

Definition constraint_semantic (c: constraint) : expr_b := (nat_e 0) >>= c.

Definition andlist := list constraint.

Fixpoint andlist_semantic (l: andlist) : expr_b:=
 match l with
   nil => true_b
   | hd::tl => (constraint_semantic hd) &&& (andlist_semantic tl)
 end.

Definition andlist_plus_andlist (c1 c2: andlist) : andlist :=
  c1 ++ c2.

Lemma andlist_plus_andlist_sem: forall b1 b2,
  (forall s,
    eval_b (andlist_semantic (andlist_plus_andlist b1 b2)) s =
    eval_b (andlist_semantic b1) s && eval_b (andlist_semantic b2) s
  ).

Lemma andlist_semantic_app : forall l1 l2 s,
 eval_b (andlist_semantic (l1 ++ l2)) s =
 eval_b ((andlist_semantic l1) &&& (andlist_semantic l2)) s.

Definition orlist := list andlist.

Fixpoint orlist_semantic (l: orlist) : expr_b:=
 match l with
   nil => ! true_b
   | hd::tl => (andlist_semantic hd) ||| (orlist_semantic tl)
 end.

Definition orlist_plus_orlist (d1 d2: orlist) : orlist :=
  d1 ++ d2.

Lemma orlist_plus_orlist_sem: forall b1 b2,
  (forall s,
    eval_b (orlist_semantic (orlist_plus_orlist b1 b2)) s =
    eval_b (orlist_semantic b1) s || eval_b (orlist_semantic b2) s
  ).

Fixpoint andlist_mult_orlist (c: andlist) (d: orlist) {struct d} : orlist :=
  match d with
    | nil => nil
    | hd::tl => orlist_plus_orlist ((andlist_plus_andlist c hd)::nil) (andlist_mult_orlist c tl)
  end.

Lemma andlist_mult_orlist_sem: forall b2 b1,
  (forall s,
    eval_b (orlist_semantic (andlist_mult_orlist b1 b2)) s =
    eval_b (andlist_semantic b1) s &&
    eval_b (orlist_semantic b2) s
  ).

Fixpoint orlist_mult_orlist (d1 d2: orlist) {struct d1} : orlist :=
  match d1 with
    | nil => nil
    | hd::tl => orlist_plus_orlist (andlist_mult_orlist hd d2) (orlist_mult_orlist tl d2)
  end.

Lemma orlist_mult_orlist_sem: forall b1 b2,
  (forall s,
    eval_b (orlist_semantic (orlist_mult_orlist b1 b2)) s =
    eval_b (orlist_semantic b1) s &&
    eval_b (orlist_semantic b2) s
  ).

Lemma orlist_semantic_app: forall l1 l2 s,
  l1 <> nil ->
  l2 <> nil ->
  eval_b (orlist_semantic (l1 ++ l2)) s =
  eval_b ((orlist_semantic l1) ||| (orlist_semantic l2)) s.
  Eval_b_hyp_clean; Eval_b_goal.

Fixpoint disj_nf (b: expr_b) : orlist :=
  match b with
    | and_b e1 e2 =>
      orlist_mult_orlist (disj_nf e1) (disj_nf e2)
    | or_b e1 e2 =>
      orlist_plus_orlist (disj_nf e1) (disj_nf e2)
    | neg_b b1 =>
      match b1 with
        | true_b => ((nat_e 1)::nil)::nil
        | eq_b e1 e2 => ((e1 +e (nat_e 1) -e e2)::nil)::((e2 +e (nat_e 1) -e e1)::nil)::nil
        | neq_b e1 e2 => ((e1 -e e2)::(e2 -e e1)::nil)::nil
        | ge_b e1 e2 => ((e1 +e (nat_e 1) -e e2)::nil)::nil
        | gt_b e1 e2 => ((e1 -e e2)::nil)::nil
        | _ => nil
      end
    | true_b => ((nat_e 0)::nil)::nil
    | eq_b e1 e2 => ((e1 -e e2)::(e2 -e e1)::nil)::nil
    | neq_b e1 e2 => ((e1 +e (nat_e 1) -e e2)::nil)::((e2 +e (nat_e 1) -e e1)::nil)::nil
    | ge_b e1 e2 => ((e2 -e e1)::nil)::nil
    | gt_b e1 e2 => ((e2 +e (nat_e 1) -e e1)::nil)::nil
  end.

Lemma disj_nf_preserve: forall b,
  is_neg_propagate b ->
  (forall s, eval_b (orlist_semantic (disj_nf b)) s = eval_b b s).

Definition expr_b2constraints (b: expr_b) : orlist :=
  disj_nf (neg_propagate b false).

Lemma expr_b2constraints_correct: forall b,
  (forall s, eval_b (orlist_semantic (expr_b2constraints b)) s = eval_b b s).

Open Local Scope Z_scope.

Fixpoint expr_compute (e: expr) : option Z :=
   match e with
     var_e x => None
     | int_e x => Some x
     | e1 +e e2 => match expr_compute e1 with
                     None => None
                     | Some e1' =>
                       match expr_compute e2 with
                         None => None
                         | Some e2' => Some (e1' + e2')
                       end
                   end
     | e1 -e e2 => match expr_compute e1 with
                     None => None
                     | Some e1' =>
                       match expr_compute e2 with
                         None => None
                         | Some e2' => Some (e1' - e2')
                       end
                   end
     | e1 *e e2 => match expr_compute e1 with
                     None =>
                     match expr_compute e2 with
                       None => None
                       | Some e2' => if Zeq_bool e2' 0 then Some 0 else None
                     end
                     | Some e1' =>
                       if Zeq_bool e1' 0 then Some 0 else
                         match expr_compute e2 with
                           None => None
                           | Some e2' => Some (e1' * e2')
                         end
                   end
     | _ => None
   end.

Lemma expr_compute_correct: forall e z,
  expr_compute e = Some z ->
  (forall s, eval e s = z).

Fixpoint simpl_expr' (e: expr) : expr :=
  match e with
    | var_e v => var_e v
    | int_e z => int_e z
    | e1 +e e2 =>
      match expr_compute (simpl_expr' e1) with
        | None =>
          match expr_compute (simpl_expr' e2) with
            | None => e
            | Some e2' =>
              if Zeq_bool e2' 0 then e1 else e1 +e (int_e e2')
          end
        | Some e1' =>
          if Zeq_bool e1' 0 then
            (match expr_compute (simpl_expr' e2) with
                | None => e2
                | Some e2' =>
                  if Zeq_bool e2' 0 then int_e 0 else int_e e2'
              end
            ) else (
              match expr_compute (simpl_expr' e2) with
                | None => (int_e e1') +e e2
                | Some e2' =>
                  if Zeq_bool e2' 0 then int_e e1' else int_e (e1' + e2')
              end)
      end
    | e1 -e e2 =>
      match expr_compute (simpl_expr' e1) with
        | None =>
          match expr_compute (simpl_expr' e2) with
            | None => e
            | Some e2' =>
              if Zeq_bool e2' 0 then e1 else e1 -e (int_e e2')
          end
        | Some e1' =>
          if Zeq_bool e1' 0 then
            (match expr_compute (simpl_expr' e2) with
                | None => (int_e 0 -e e2)
                | Some e2' =>
                  if Zeq_bool e2' 0 then int_e 0 else int_e ( - e2')
              end
            ) else (
              match expr_compute (simpl_expr' e2) with
                | None => (int_e e1') -e e2
                | Some e2' =>
                  if Zeq_bool e2' 0 then int_e e1' else int_e (e1' - e2')
              end)
      end
    | e1 *e e2 =>
      match expr_compute (simpl_expr' e1) with
        | None =>
          match expr_compute (simpl_expr' e2) with
            | None => e
            | Some e2' =>
              if Zeq_bool e2' 0 then
                int_e 0 else
                  if Zeq_bool e2' 1 then e1 else e1 *e (int_e e2')
          end
        | Some e1' =>
          if Zeq_bool e1' 0 then
            int_e 0 else
              if Zeq_bool e1' 1 then
                (match expr_compute (simpl_expr' e2) with
                    | None => e2
                    | Some e2' =>
                      if Zeq_bool e2' 0 then int_e 0 else int_e (e2')
                  end
                ) else
                (match expr_compute (simpl_expr' e2) with
                    | None => (int_e e1') *e e2
                    | Some e2' =>
                      if Zeq_bool e2' 0 then
                        int_e 0 else
                          if Zeq_bool e2' 1 then int_e e1' else int_e (e1' * e2')
                  end)
      end
    | _ => e
  end.

Lemma simpl_expr'_correct: forall e s,
  eval e s = eval (simpl_expr' e) s.

Require Import Max.

Open Local Scope nat_scope.

Fixpoint expr_deep (e: expr) : nat :=
  match e with
    | e1 +e e2 => 1 + (max (expr_deep e1) (expr_deep e2))
    | e1 -e e2 => 1 + (max (expr_deep e1) (expr_deep e2))
    | e1 *e e2 => 1 + (max (expr_deep e1) (expr_deep e2))
    | _ => 1
  end.

Close Local Scope nat_scope.

Fixpoint simpl_expr_fp (e: expr) (n: nat) {struct n} : expr :=
  match n with
    | O => e
    | S n' =>
      match e with
        | var_e v => var_e v
        | int_e z => int_e z
        | e1 +e e2 => simpl_expr' ((simpl_expr_fp e1 n') +e (simpl_expr_fp e2 n'))
        | e1 -e e2 => simpl_expr' ((simpl_expr_fp e1 n') -e (simpl_expr_fp e2 n'))
        | e1 *e e2 => simpl_expr' ((simpl_expr_fp e1 n') *e (simpl_expr_fp e2 n'))
        | _ => e
      end
  end.

Opaque simpl_expr'.

Lemma simpl_expr_fp_corect: forall n e s,
  eval e s = eval (simpl_expr_fp e n) s.

Definition simpl_expr (e: expr) : expr :=
  simpl_expr_fp e (expr_deep e).

Lemma simpl_expr_corect: forall e s,
  eval e s = eval (simpl_expr e) s.

Transparent simpl_expr'.

Fixpoint expr_var_fact' (e:expr) (v:nat) {struct e} : (expr * expr) :=
  match e with
    var_e x => if beq_nat x v then (nat_e 1, nat_e 0) else (nat_e 0, var_e x)
    | int_e x => (nat_e 0, int_e x)
    | e1 +e e2 => match expr_var_fact' e1 v with
                    | (e11, e12) => match expr_var_fact' e2 v with
                                      | (e21,e22) => (e11 +e e21, e12 +e e22)
                                    end
                  end
    | e1 -e e2 => match expr_var_fact' e1 v with
                    | (e11, e12) => match expr_var_fact' e2 v with
                                     | (e21, e22) => (e11 -e e21, e12 -e e22)
                                   end
                  end
    | e1 *e e2 => match expr_var_fact' e1 v with
                    | (e11, e12) => match expr_var_fact' e2 v with
                                     | (e21, e22) =>
                                       (((e12 *e e21) +e (e11 *e e22))
                                         +e
                                         (var_e v *e (e11 *e e21)),
                                       e12 *e e22)
                                   end
                  end
    | e1 /e e2 => (nat_e 0, e1 /e e2)
    | e1 mode e2 => (nat_e 0, e1 mode e2)
  end.

Lemma expr_var_fact'_correct: forall e v e1 e2,
  expr_var_fact' e v = (e1, e2) ->
  (forall s, eval e s = eval ((var_e v *e e1) +e e2) s).

Fixpoint simpl_varlist_constraint (c:constraint) (v:list nat) {struct v} : constraint :=
  match v with
    nil => simpl_expr c
    | hd :: tl =>
      match expr_var_fact' c hd with
        | (e1, e2) =>
          simpl_expr (
            (simpl_expr (var_e hd *e (simpl_varlist_constraint e1 tl)))
            +e
            (simpl_expr (simpl_varlist_constraint e2 tl)) )
      end
  end.

Lemma simpl_varlist_constraint_correct: forall v c c',
 simpl_varlist_constraint c v = c' ->
 (forall s, eval c s = eval c' s).

Definition simpl_constraint (c:constraint) : constraint :=
  simpl_varlist_constraint c (Expr_var c).

Lemma simpl_constraint_correct: forall c,
  (forall s, eval c s = eval (simpl_constraint c) s).

Fixpoint simpl_expr_b (b:expr_b) : expr_b :=
  match b with
    | eq_b e1 e2 => eq_b (simpl_constraint e1) (simpl_constraint e2)
    | neq_b e1 e2 => neq_b (simpl_constraint e1) (simpl_constraint e2)
    | ge_b e1 e2 => ge_b (simpl_constraint e1) (simpl_constraint e2)
    | gt_b e1 e2 => gt_b (simpl_constraint e1) (simpl_constraint e2)
    | and_b e1 e2 => and_b (simpl_expr_b e1) (simpl_expr_b e2)
    | or_b e1 e2 => or_b (simpl_expr_b e1) (simpl_expr_b e2)
    | neg_b e => neg_b (simpl_expr_b e)
    | _ => b
  end.

Lemma simpl_expr_b_correct: forall b s,
  eval_b b s = eval_b (simpl_expr_b b) s.

Fixpoint simpl_andlist (l:andlist) : andlist :=
  match l with
    nil => nil
    | hd :: tl => simpl_constraint hd :: simpl_andlist tl
  end.

Lemma simpl_andlist_correct: forall l s,
  eval_b (andlist_semantic l) s = eval_b (andlist_semantic (simpl_andlist l)) s.

Fixpoint simpl_orlist (l:orlist) : orlist :=
  match l with
    nil => nil
    | hd :: tl => simpl_andlist hd :: simpl_orlist tl
  end.

Lemma simpl_orlist_correct: forall l s,
  eval_b (orlist_semantic l) s = eval_b (orlist_semantic (simpl_orlist l)) s.
  Eval_b_hyp_clean.
  Eval_b_hyp_clean.

Definition expr_var_fact (e:expr) (v:nat) : expr * expr :=
   match expr_var_fact' e v with
     (e1, e2) => (simpl_constraint e1, simpl_constraint e2)
   end.

Lemma expr_var_fact_correct: forall c v e1 e2,
  expr_var_fact c v = (e1, e2) ->
  (forall s, eval c s = eval ((var_e v *e e1) +e e2) s).

Definition elim_var_constraint (c1 c2:constraint) (v:nat) : constraint :=
  match expr_var_fact c1 v with
    | (e11, e12) =>
      match expr_var_fact c2 v with
        | (e21, e22) =>
          match expr_compute (simpl_constraint e11) with
            | None => simpl_constraint c2
            | Some e11' =>
              match expr_compute (simpl_constraint e21) with
                | None => simpl_constraint c2
                | Some e21' =>
                  if andb (Zlt_bool e11' 0) (Zlt_bool 0 e21') then
                    simpl_constraint ((e21 *e e12) -e (e11 *e e22))
                    else
                      (if andb (Zlt_bool 0 e11') (Zlt_bool e21' 0) then
                        simpl_constraint ((e11 *e e22) -e (e21 *e e12))
                        else
                          simpl_constraint c2)
              end
          end
      end
   end.

Lemma fourier_motzkin_for_integers: forall a1 b1 a2 b2 x,
 a1 < 0 ->
 0 < a2 ->
 0 >= x * a1 + b1 ->
 0 >= x * a2 + b2 ->
 a1 * b2 >= a2 * b1.

Lemma expr_b_semantic_good' : forall e s,
  eval_b e s = true -> expr_b_sem e s.

Lemma elim_var_constraint_correct: forall c1 c2 v s,
  ((s |b= (constraint_semantic c1) &&& (constraint_semantic c2)) ->
    (s |b= (constraint_semantic (elim_var_constraint c1 c2 v)))).

Fixpoint elim_var_andlist' (c: constraint) (l l':andlist) (v: nat) {struct l} : andlist :=
  match l with
    nil => l'
    | hd :: tl => if inb beq_nat (Expr_var (simpl_constraint hd)) v then
      elim_var_andlist' c tl ((elim_var_constraint c hd v)::l') v
      else
        elim_var_andlist' c tl l' v
  end.

Lemma elim_var_andlist'_correct: forall l l' c v,
  (forall s, (s |b= (andlist_semantic (c::l ++ l'))) ->
                s |b= (andlist_semantic (elim_var_andlist' c l l' v))).

Fixpoint elim_var_andlist (l l': andlist) (v: nat) {struct l} : andlist :=
  match l with
    nil => l'
    | hd::tl => if inb beq_nat (Expr_var (simpl_constraint hd)) v then
      elim_var_andlist tl (l' ++ (elim_var_andlist' hd tl nil v)) v
      else
        elim_var_andlist tl (hd::l') v
  end.

Lemma elim_var_andlist_correct: forall l l' v,
  (forall s, (s |b= (andlist_semantic (l ++ l'))) ->
                s |b= (andlist_semantic (elim_var_andlist l l' v))).

Fixpoint elim_var_orlist (l: orlist) (v: nat) {struct l} : orlist :=
   match l with
      nil => nil
      | hd::tl => (elim_var_andlist hd nil v)::(elim_var_orlist tl v)
   end.

Lemma elim_var_orlist_correct: forall l v,
  (forall s, eval_b (orlist_semantic l) s = true ->
               eval_b (orlist_semantic (elim_var_orlist l v)) s = true ).
Eval_b_hyp_clean.

Fixpoint elim_varlist_orlist (l: orlist) (v: list nat) {struct v} : orlist :=
   match v with
      nil => simpl_orlist l
      | hd::tl => (elim_varlist_orlist (elim_var_orlist l hd) tl)
   end.

Lemma elim_varlist_orlist_correct: forall v l,
 (forall s, eval_b (orlist_semantic l) s = true ->
               eval_b (orlist_semantic (elim_varlist_orlist l v)) s = true).

Definition eval_constraint (c : constraint) : option bool :=
   match expr_compute c with
     | Some z => Some (Zge_bool 0 z)
     | _ => None
   end.

Lemma eval_constraint2constraint_semantic: forall c b,
  eval_constraint c = Some b ->
  (forall s, eval_b (constraint_semantic c) s = b).

Fixpoint eval_andlist' (a: andlist) : option bool :=
  match a with
    nil => Some true
    | hd :: tl => match eval_constraint hd with
                  | Some false => Some false
                    | o => match eval_andlist' tl with
                             | None => None
                             | Some false => Some false
                             | Some true => andb_option o (Some true)
                           end
                end
  end.

Lemma eval_andlist'2andlist_semantic: forall a b,
  eval_andlist' a = Some b ->
  (forall s, eval_b (andlist_semantic a) s = b).



Definition eval_andlist (a: andlist) : option bool :=
  if beq_nat (length a) 0 then
      None
  else
      eval_andlist' a.

Lemma eval_andlist2andlist_semantic: forall a b,
  eval_andlist a = Some b ->
  (forall s, eval_b (andlist_semantic a) s = b).

Fixpoint eval_orlist' (o: orlist) : option bool :=
   match o with
     nil => Some false
     | hd :: tl => orb_option (eval_andlist hd) (eval_orlist' tl)
   end.

Lemma eval_orlist'2orlist_semantic: forall a b,
  eval_orlist' a = Some b ->
  (forall s, eval_b (orlist_semantic a) s = b).

Definition eval_orlist (a: orlist) : option bool :=
  if beq_nat (length a) 0 then
    None
    else
      eval_orlist' a.

Lemma eval_orlist2orlist_semantic: forall a b,
  eval_orlist a = Some b ->
  (forall s, eval_b (orlist_semantic a) s = b).

Definition expr_b_dp (b: expr_b) : bool :=
  match eval_orlist (elim_varlist_orlist (simpl_orlist (expr_b2constraints (simpl_expr_b (!b)))) (Expr_B_var (simpl_expr_b b))) with
    | None => false
    | Some res => negb res
  end.

Lemma expr_b_dp_correct: forall b,
   expr_b_dp b = true ->
   (forall s, s |b= b).

Lemma expr_b_impl_intro: forall b1 b2 s,
  (s |b= b1) ->
  (s |b= (b1 =b> b2)) ->
  (s |b= b2).

Ltac Is_pos_var p :=
  match p with
    xH => false
    | xI ?n => false
    | xO ?n => false
    | _ => true
  end.

Ltac Is_Z_var t :=
  match t with
    | Z0 => false
    | Zpos ?e => Is_pos_var e
    | Zneg ?e => Is_pos_var e
    | _ => true
  end.

Ltac Add_list e l :=
  match l with
    | e::?tl => l
    | ?hd::?tl => let x := Add_list e tl in ( constr:(hd::x) )
    | (@nil Z) => constr:(e::nil)
  end.

Ltac Concat_list l1 l2 :=
  match l1 with
    | ?hd::?tl => let x:= Add_list hd l2 in ( Concat_list tl x )
    | nil => l2
  end.

Ltac Build_env t :=
  match t with
    | (?t1 + ?t2)%Z =>
      let x:= (Build_env t1) in (
        let y:= (Build_env t2) in (Concat_list x y)
      )
    | (?t1 - ?t2)%Z =>
      let x:= (Build_env t1) in (
        let y:= (Build_env t2) in (Concat_list x y)
      )
    | (?t1 * ?t2)%Z =>
      let x:= (Build_env t1) in (
        let y:= (Build_env t2) in (Concat_list x y)
      )
    | (?t1 = ?t2)%Z =>
      let x:= (Build_env t1) in (
        let y:= (Build_env t2) in (Concat_list x y)
      )
    | (?t1 -> ?t2)%Z =>
      let x:= (Build_env t1) in (
        let y:= (Build_env t2) in (Concat_list x y)
      )
    | (?t1 > ?t2)%Z =>
      let x:= (Build_env t1) in (
        let y:= (Build_env t2) in (Concat_list x y)
      )
    | (?t1 < ?t2)%Z =>
      let x:= (Build_env t1) in (
        let y:= (Build_env t2) in (Concat_list x y)
      )
    | (?t1 >= ?t2)%Z =>
      let x:= (Build_env t1) in (
        let y:= (Build_env t2) in (Concat_list x y)
      )
    | (?t1 <= ?t2)%Z =>
      let x:= (Build_env t1) in (
        let y:= (Build_env t2) in (Concat_list x y)
      )
    | (?t1 /\ ?t2)%Z =>
      let x:= (Build_env t1) in (
        let y:= (Build_env t2) in (Concat_list x y)
      )
    | (?t1 \/ ?t2)%Z =>
      let x:= (Build_env t1) in (
        let y:= (Build_env t2) in (Concat_list x y)
      )
    | (?t1 <> ?t2)%Z =>
      let x:= (Build_env t1) in (
        let y:= (Build_env t2) in (Concat_list x y)
      )
    | (~ ?t1)%Z =>
      let x:= (Build_env t1) in (x)
    | _ =>
      let x:=(Is_Z_var t) in (
        match (eval compute in x) with
          true => constr:(t::nil)
          | false => constr:(@nil Z)
        end
        )
  end.

Ltac Find_var v l c :=
  match l with
    | v::?tl => constr:c
    | ?hd::?tl => let x := eval compute in (S c) in Find_var v tl x
    | (@nil Z) => constr:0
  end.

Ltac Get_var_index v l := Find_var v l O.

Ltac Print t := assert (t = t).

Ltac To_expr t l :=
  match t with
    | (?t1 + ?t2)%Z =>
      let x:= (To_expr t1 l) in (
        let y:= (To_expr t2 l) in (
          constr:(x +e y)
        )
      )
    | (?t1 - ?t2)%Z =>
      let x:= (To_expr t1 l) in (
        let y:= (To_expr t2 l) in (
          constr:(x -e y)
        )
      )
    | (?t1 * ?t2)%Z =>
      let x:= (To_expr t1 l) in (
        let y:= (To_expr t2 l) in (
          constr:(x *e y)
        )
      )
    | _ =>
      let x:= Is_Z_var t in (
        match eval compute in x with
          | true =>
            let y:= (Get_var_index t l) in
             constr:(var_e y)
          | false =>
             constr:(int_e t)
        end
        )
  end.

Ltac To_expr_b t l :=
  match t with
    | (?t1 = ?t2) =>
      let x := To_expr t1 l in (
        let y := To_expr t2 l in (constr:(x == y))
      )
    | (?t1 > ?t2)%Z =>
      let x := To_expr t1 l in (
        let y := To_expr t2 l in (constr:(x >> y))
      )
    | (?t1 >= ?t2)%Z =>
      let x := To_expr t1 l in (
        let y := To_expr t2 l in (constr:(x >>= y))
      )
    | (?t1 < ?t2)%Z =>
      let x := To_expr t1 l in (
        let y := To_expr t2 l in (constr:(! (x >>= y)))       )
    | (?t1 <= ?t2)%Z =>
      let x := To_expr t1 l in (
        let y := To_expr t2 l in (constr:(! (x >> y)))       )
    | (?t1 -> ?t2) =>
      let x := To_expr_b t1 l in (
        let y := To_expr_b t2 l in (constr:(x =b> y))
      )
    | (?t1 /\ ?t2) =>
      let x := To_expr_b t1 l in (
        let y := To_expr_b t2 l in (constr:(x &&& y))
      )
    | (?t1 \/ ?t2) =>
      let x := To_expr_b t1 l in (
        let y := To_expr_b t2 l in (constr:(x ||| y))
      )
    | (?t1 <> ?t2) =>
      let x := To_expr_b t1 l in (
        let y := To_expr_b t2 l in (constr:(x =/= y))
      )
    | (~ ?t1) =>
      let x := To_expr_b t1 l in (constr:(! x))
  end.

Ltac In_list v l :=
  match l with
    | v::?tl => true
    | ?hd::?tl => In_list v tl
    | (@nil Z) => false
  end.

Ltac In_list_list l1 l2 :=
  match l1 with
    | ?hd::?tl =>
      let x := In_list hd l2 in
        match eval compute in x with
          | true => true
          | false => (In_list_list tl l2)
        end
    | (@nil Z) =>
      false
  end.

Lemma new_cut: forall P Q,
  P -> (P -> Q) -> Q.

Ltac new_cut P :=
  match goal with
    | |- ?Q =>
      apply (new_cut P Q)
  end.

Fixpoint list2store (l: list Z) {struct l} : store.s :=
   match l with
     | nil => store.emp
     | hd::tl => store.update (length l - 1)%nat hd (list2store tl)
   end.

Lemma lookup_list2store: forall l x,
   store.lookup x (list2store l) = nth x (rev l) 0%Z.

Opaque list2store.

Ltac expr_b_dp_decision :=
    match goal with
    | |- ?G =>
            let l := (Build_env G) in (
                let x := (To_expr_b G l) in (
                    new_cut (eval_b x (list2store (rev l)) = true); [
                        eapply expr_b_dp_correct; compute; apply refl_equal
              |
                          let h:= fresh in (
                intro h; let y := fresh in (
                  generalize (proj1 (expr_b_semantic_good x (list2store (rev l))) h); intro y; simpl in y; repeat (rewrite lookup_list2store in y); simpl in y; firstorder
                )
             )
          ]
        )
      )
  end.

Ltac Intro_hyp_var_list l :=
  match goal with
    | id: ?a%Z |- _ =>
      let l' := (Build_env a) in (
        let x := In_list_list l' l in (
          match (eval compute in x) with
            | true => generalize id; clear id; (Expr_b_dp || expr_b_dp_decision)
            | false => fail
          end
        )
      )

    | _ => expr_b_dp_decision
  end
with Expr_b_dp :=
  match goal with
    | |- ?G =>
      let l := (Build_env G) in (
        Intro_hyp_var_list l
      )
  end.

Ltac Contradiction_b :=
  match goal with
    | id: ?A |- _ =>
      assert (~A); [Expr_b_dp | tauto]
    | _ => fail
  end.