Require terminaison.

Require Relations.

Require term.

Require List.

Require equational_theory.

Require rpo_extension.

Require equational_extension.

Require closure_extension.

Require term_extension.

Require dp.

Require Inclusion.

Require or_ext_generated.

Require ZArith.

Require ring_extention.

Require Zwf.

Require Inverse_Image.

Require matrix.

Require more_list_extention.

Import List.

Import ZArith.

Set Implicit Arguments.

Module algebra.
 Module F
  <:term.Signature.
  Inductive symb  : Set := 
     (* id__dot_ *)
    | id__dot_ : symb
  .
  
  
  Definition symb_eq_bool (f1 f2:symb) : bool := 
    match f1,f2 with
      | id__dot_,id__dot_ => true
      end.
  
  
   (* Proof of decidability of equality over symb *)
  Definition symb_eq_bool_ok(f1 f2:symb) :
   match symb_eq_bool f1 f2 with
     | true => f1 = f2
     | false => f1 <> f2
     end.
  Proof.
    intros f1 f2.
    
    refine match f1 as u1,f2 as u2 return 
             match symb_eq_bool u1 u2 return 
               Prop with
               | true => u1 = u2
               | false => u1 <> u2
               end with
             | id__dot_,id__dot_ => refl_equal _
             end;intros abs;discriminate.
  Defined.
  
  Definition arity (f:symb) := match f with
                                 | id__dot_ => term.Free 2
                                 end.
  
  
  Definition symb_order (f1 f2:symb) : bool := 
    match f1,f2 with
      | id__dot_,id__dot_ => true
      end.
  
  
  Module Symb.
   Definition A  := symb.
   
   Definition eq_A  := @eq A.
   
   
   Definition eq_proof : equivalence A eq_A.
   Proof.
     constructor.
     red ;reflexivity .
     red ;intros ;transitivity y ;assumption.
     red ;intros ;symmetry ;assumption.
   Defined.
   
   
   Add Relation A eq_A 
  reflexivity proved by (@equiv_refl _ _ eq_proof)
    symmetry proved by (@equiv_sym _ _ eq_proof)
      transitivity proved by (@equiv_trans _ _ eq_proof) as EQA
.
   
   Definition eq_bool  := symb_eq_bool.
   
   Definition eq_bool_ok  := symb_eq_bool_ok.
  End Symb.
  
  Export Symb.
 End F.
 
 Module Alg := term.Make'(F)(term_extension.IntVars).
 
 Module Alg_ext := term_extension.Make(Alg).
 
 Module EQT := equational_theory.Make(Alg).
 
 Module EQT_ext := equational_extension.Make(EQT).
End algebra.

Module R_xml_0_deep_rew.
 Inductive R_xml_0_rules  :
  algebra.Alg.term ->algebra.Alg.term ->Prop := 
    (* .(.(x_,y_),z_) -> .(x_,.(y_,z_)) *)
   | R_xml_0_rule_0 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id__dot_ ((algebra.Alg.Var 1)::
                   (algebra.Alg.Term algebra.F.id__dot_ 
                   ((algebra.Alg.Var 2)::(algebra.Alg.Var 3)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id__dot_ ((algebra.Alg.Term 
      algebra.F.id__dot_ ((algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil))::
      (algebra.Alg.Var 3)::nil))
 .
 
 
 Definition R_xml_0_rule_as_list_0  := 
   ((algebra.Alg.Term algebra.F.id__dot_ ((algebra.Alg.Term 
     algebra.F.id__dot_ ((algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil))::
     (algebra.Alg.Var 3)::nil)),
    (algebra.Alg.Term algebra.F.id__dot_ ((algebra.Alg.Var 1)::
     (algebra.Alg.Term algebra.F.id__dot_ ((algebra.Alg.Var 2)::
     (algebra.Alg.Var 3)::nil))::nil)))::nil.
 
 Definition R_xml_0_rule_as_list  := R_xml_0_rule_as_list_0.
 
 
 Lemma R_xml_0_rules_included :
  forall l r, R_xml_0_rules r l <-> In (l,r) R_xml_0_rule_as_list.
 Proof.
   intros l r.
   constructor.
   intros H.
   
   case H;clear H;
    (apply (more_list.mem_impl_in (@eq (algebra.Alg.term*algebra.Alg.term)));
     [tauto|idtac]);
    match goal with
      |  |- _ _ _ ?t ?l =>
       let u := fresh "u" in 
        (generalize (more_list.mem_bool_ok _ _ 
                      algebra.Alg_ext.eq_term_term_bool_ok t l);
          set (u:=more_list.mem_bool algebra.Alg_ext.eq_term_term_bool t l) in *;
          vm_compute in u|-;unfold u in *;clear u;intros H;refine H)
      end
    .
   intros H.
   case H;clear H;intros H.
   injection H;intros ;subst;constructor 1.
   elim H.
 Qed.
 
 
 Lemma R_xml_0_non_var : forall x t, ~R_xml_0_rules t (algebra.EQT.T.Var x).
 Proof.
   intros x t H.
   inversion H.
 Qed.
 
 
 Lemma R_xml_0_reg :
  forall s t, 
   (R_xml_0_rules s t) ->
    forall x, In x (algebra.Alg.var_list s) ->In x (algebra.Alg.var_list t).
 Proof.
   intros s t H.
   
   inversion H;intros x Hx;
    (apply (more_list.mem_impl_in (@eq algebra.Alg.variable));[tauto|idtac]);
    apply (more_list.in_impl_mem (@eq algebra.Alg.variable)) in Hx;
    vm_compute in Hx|-*;tauto.
 Qed.
 
 Ltac impossible_star_reduction_R_xml_0  :=
  fail.
 
 Ltac simplify_star_reduction_R_xml_0  :=
  fail.
End R_xml_0_deep_rew.

Module InterpGen := interp.Interp(algebra.EQT).

Module ddp := dp.MakeDP(algebra.EQT).

Module SymbType. Definition A := algebra.Alg.F.Symb.A. End SymbType.

Module Symb_more_list := more_list_extention.Make(SymbType)(algebra.Alg.F.Symb).

Module SymbSet := list_set.Make(algebra.F.Symb).

Module Interp.
 Section S.
   Require Import interp.
   
   Hypothesis A : Type.
   
   Hypothesis Ale Alt Aeq : A -> A -> Prop.
   
   Hypothesis Aop : interp.ordering_pair Aeq Alt Ale.
   
   Hypothesis A0 : A.
   
   Notation Local "a <= b" := (Ale a b).
   
   Hypothesis P_id__dot_ : A ->A ->A.
   
   Hypothesis P_id__dot__monotonic :
    forall x8 x6 x5 x7, 
     (A0 <= x8)/\ (x8 <= x7) ->
      (A0 <= x6)/\ (x6 <= x5) ->P_id__dot_ x6 x8 <= P_id__dot_ x5 x7.
   
   Hypothesis P_id__dot__bounded :
    forall x6 x5, (A0 <= x5) ->(A0 <= x6) ->A0 <= P_id__dot_ x6 x5.
   
   Fixpoint measure t { struct t }  := 
     match t with
       | (algebra.Alg.Term algebra.F.id__dot_ (x6::x5::nil)) =>
        P_id__dot_ (measure x6) (measure x5)
       | _ => A0
       end.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id__dot_ (x6::x5::nil)) =>
                    P_id__dot_ (measure x6) (measure x5)
                   | _ => A0
                   end.
   Proof.
     intros t;case t;intros ;apply refl_equal.
   Qed.
   
   Definition Pols f : InterpGen.Pol_type A (InterpGen.get_arity f) := 
     match f with
       | algebra.F.id__dot_ => P_id__dot_
       end.
   
   Lemma same_measure : forall t, measure t = InterpGen.measure A0 Pols t.
   Proof.
     fix 1 .
     intros [a| f l].
     simpl in |-*.
     unfold eq_rect_r, eq_rect, sym_eq in |-*.
     reflexivity .
     
     refine match f with
              | algebra.F.id__dot_ =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              end;simpl in |-*;unfold eq_rect_r, eq_rect, sym_eq in |-*;
      try (reflexivity );f_equal ;auto.
   Qed.
   
   Lemma measure_bounded : forall t, A0 <= measure t.
   Proof.
     intros t.
     rewrite same_measure in |-*.
     apply (InterpGen.measure_bounded Aop).
     intros f.
     case f.
     vm_compute in |-*;intros ;apply P_id__dot__bounded;assumption.
   Qed.
   
   Ltac generate_pos_hyp  :=
    match goal with
      | H:context [measure ?x] |- _ =>
       let v := fresh "v" in 
        (let H := fresh "h" in 
          (set (H:=measure_bounded x) in *;set (v:=measure x) in *;
            clearbody H;clearbody v))
      |  |- context [measure ?x] =>
       let v := fresh "v" in 
        (let H := fresh "h" in 
          (set (H:=measure_bounded x) in *;set (v:=measure x) in *;
            clearbody H;clearbody v))
      end
    .
   
   Hypothesis rules_monotonic :
    forall l r, 
     (algebra.EQT.axiom R_xml_0_deep_rew.R_xml_0_rules r l) ->
      measure r <= measure l.
   
   Lemma measure_star_monotonic :
    forall l r, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                r l) ->measure r <= measure l.
   Proof.
     intros .
     do 2 (rewrite same_measure in |-*).
     
     apply InterpGen.measure_star_monotonic with (1:=Aop) (Pols:=Pols) 
     (rules:=R_xml_0_deep_rew.R_xml_0_rules).
     intros f.
     case f.
     vm_compute in |-*;intros ;apply P_id__dot__monotonic;assumption.
     intros f.
     case f.
     vm_compute in |-*;intros ;apply P_id__dot__bounded;assumption.
     intros .
     do 2 (rewrite  <- same_measure in |-*).
     apply rules_monotonic;assumption.
     assumption.
   Qed.
   
   Hypothesis P_id__dot__hat_1 : A ->A ->A.
   
   Hypothesis P_id__dot__hat_1_monotonic :
    forall x8 x6 x5 x7, 
     (A0 <= x8)/\ (x8 <= x7) ->
      (A0 <= x6)/\ (x6 <= x5) ->
       P_id__dot__hat_1 x6 x8 <= P_id__dot__hat_1 x5 x7.
   
   Definition marked_measure t := 
     match t with
       | (algebra.Alg.Term algebra.F.id__dot_ (x6::x5::nil)) =>
        P_id__dot__hat_1 (measure x6) (measure x5)
       | _ => measure t
       end.
   
   Definition  Marked_pols :
    forall f, 
     (algebra.EQT.defined R_xml_0_deep_rew.R_xml_0_rules f) ->
      InterpGen.Pol_type A (InterpGen.get_arity f).
   Proof.
     intros f H.
     
     apply ddp.defined_list_complete with (1:=R_xml_0_deep_rew.R_xml_0_rules_included) in H .
     apply (Symb_more_list.change_in algebra.F.symb_order) in H .
     
     set (u := (Symb_more_list.qs algebra.F.symb_order
           (Symb_more_list.XSet.remove_red
              (ddp.defined_list R_xml_0_deep_rew.R_xml_0_rule_as_list)))) in * .
     vm_compute in u .
     unfold u in * .
     clear u .
     unfold more_list.mem_bool in H .
     
     match type of H with 
       | orb ?a ?b = true => 
         assert (H':{a = true}+{b = true});[
           revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]|
             clear H;destruct H' as [H|H]]
     end .
     
     match type of H with 
       | _ _ ?t = true =>
         generalize (algebra.F.symb_eq_bool_ok f t);
           unfold algebra.Alg.eq_symb_bool in H;
           rewrite H;clear H;intros ;subst
     end .
     vm_compute in |-*;intros x6 x5;apply (P_id__dot__hat_1 x6 x5).
     discriminate H.
   Defined.
   
   Lemma same_marked_measure :
    forall t, 
     marked_measure t = InterpGen.marked_measure A0 Pols Marked_pols 
                         (ddp.defined_dec _ _ 
                           R_xml_0_deep_rew.R_xml_0_rules_included) t.
   Proof.
     intros [a| f l].
     simpl in |-*.
     unfold eq_rect_r, eq_rect, sym_eq in |-*.
     reflexivity .
     
     refine match f with
              | algebra.F.id__dot_ =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              end.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
   Qed.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id__dot_ (x6::
                             x5::nil)) =>
                           P_id__dot__hat_1 (measure x6) (measure x5)
                          | _ => measure t
                          end.
   Proof.
     reflexivity .
   Qed.
   
   Lemma marked_measure_star_monotonic :
    forall f l1 l2, 
     (closure.refl_trans_clos (closure.one_step_list (algebra.EQT.one_step 
                                                       R_xml_0_deep_rew.R_xml_0_rules)
                                                      ) l1 l2) ->
      marked_measure (algebra.Alg.Term f l1) <= marked_measure (algebra.Alg.Term 
                                                                 f l2).
   Proof.
     intros .
     do 2 (rewrite same_marked_measure in |-*).
     
     apply InterpGen.marked_measure_star_monotonic with (1:=Aop) (Pols:=
     Pols) (rules:=R_xml_0_deep_rew.R_xml_0_rules).
     clear f.
     intros f.
     case f.
     vm_compute in |-*;intros ;apply P_id__dot__monotonic;assumption.
     clear f.
     intros f.
     case f.
     vm_compute in |-*;intros ;apply P_id__dot__bounded;assumption.
     intros .
     do 2 (rewrite  <- same_measure in |-*).
     apply rules_monotonic;assumption.
     clear f.
     intros f.
     clear H.
     intros H.
     generalize H.
     
     apply ddp.defined_list_complete with (1:=R_xml_0_deep_rew.R_xml_0_rules_included) in H .
     apply (Symb_more_list.change_in algebra.F.symb_order) in H .
     
     set (u := (Symb_more_list.qs algebra.F.symb_order
           (Symb_more_list.XSet.remove_red
              (ddp.defined_list R_xml_0_deep_rew.R_xml_0_rule_as_list)))) in * .
     vm_compute in u .
     unfold u in * .
     clear u .
     unfold more_list.mem_bool in H .
     
     match type of H with 
       | orb ?a ?b = true => 
         assert (H':{a = true}+{b = true});[
           revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]|
             clear H;destruct H' as [H|H]]
     end .
     
     match type of H with 
       | _ _ ?t = true =>
         generalize (algebra.F.symb_eq_bool_ok f t);
           unfold algebra.Alg.eq_symb_bool in H;
           rewrite H;clear H;intros ;subst
     end .
     vm_compute in |-*;intros ;apply P_id__dot__hat_1_monotonic;assumption.
     discriminate H.
     assumption.
   Qed.
   
   End S.
End Interp.

Module InterpZ.
 Section S.
   Open Scope Z_scope.
   
   Hypothesis min_value : Z.
   
   Import ring_extention.
   
   Notation Local "'Alt'" := (Zwf.Zwf min_value).
   
   Notation Local "'Ale'" := Zle.
   
   Notation Local "'Aeq'" := (@eq Z).
   
   Notation Local "a <= b" := (Ale a b).
   
   Notation Local "a < b" := (Alt a b).
   
   Hypothesis P_id__dot_ : Z ->Z ->Z.
   
   Hypothesis P_id__dot__monotonic :
    forall x8 x6 x5 x7, 
     (min_value <= x8)/\ (x8 <= x7) ->
      (min_value <= x6)/\ (x6 <= x5) ->P_id__dot_ x6 x8 <= P_id__dot_ x5 x7.
   
   Hypothesis P_id__dot__bounded :
    forall x6 x5, 
     (min_value <= x5) ->(min_value <= x6) ->min_value <= P_id__dot_ x6 x5.
   
   Definition measure  := Interp.measure min_value P_id__dot_.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id__dot_ (x6::x5::nil)) =>
                    P_id__dot_ (measure x6) (measure x5)
                   | _ => min_value
                   end.
   Proof.
     intros t;case t;intros ;apply refl_equal.
   Qed.
   
   Lemma measure_bounded : forall t, min_value <= measure t.
   Proof.
     unfold measure in |-*.
     
     apply Interp.measure_bounded with Alt Aeq;
      (apply interp.o_Z)||
      (cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;auto with zarith).
   Qed.
   
   Ltac generate_pos_hyp  :=
    match goal with
      | H:context [measure ?x] |- _ =>
       let v := fresh "v" in 
        (let H := fresh "h" in 
          (set (H:=measure_bounded x) in *;set (v:=measure x) in *;
            clearbody H;clearbody v))
      |  |- context [measure ?x] =>
       let v := fresh "v" in 
        (let H := fresh "h" in 
          (set (H:=measure_bounded x) in *;set (v:=measure x) in *;
            clearbody H;clearbody v))
      end
    .
   
   Hypothesis rules_monotonic :
    forall l r, 
     (algebra.EQT.axiom R_xml_0_deep_rew.R_xml_0_rules r l) ->
      measure r <= measure l.
   
   Lemma measure_star_monotonic :
    forall l r, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                r l) ->measure r <= measure l.
   Proof.
     unfold measure in *.
     apply Interp.measure_star_monotonic with Alt Aeq.
     
     (apply interp.o_Z)||
     (cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;auto with zarith).
     intros ;apply P_id__dot__monotonic;assumption.
     intros ;apply P_id__dot__bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Hypothesis P_id__dot__hat_1 : Z ->Z ->Z.
   
   Hypothesis P_id__dot__hat_1_monotonic :
    forall x8 x6 x5 x7, 
     (min_value <= x8)/\ (x8 <= x7) ->
      (min_value <= x6)/\ (x6 <= x5) ->
       P_id__dot__hat_1 x6 x8 <= P_id__dot__hat_1 x5 x7.
   
   Definition marked_measure  := 
     Interp.marked_measure min_value P_id__dot_ P_id__dot__hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id__dot_ (x6::
                             x5::nil)) =>
                           P_id__dot__hat_1 (measure x6) (measure x5)
                          | _ => measure t
                          end.
   Proof.
     reflexivity .
   Qed.
   
   Lemma marked_measure_star_monotonic :
    forall f l1 l2, 
     (closure.refl_trans_clos (closure.one_step_list (algebra.EQT.one_step 
                                                       R_xml_0_deep_rew.R_xml_0_rules)
                                                      ) l1 l2) ->
      marked_measure (algebra.Alg.Term f l1) <= marked_measure (algebra.Alg.Term 
                                                                 f l2).
   Proof.
     unfold marked_measure in *.
     apply Interp.marked_measure_star_monotonic with Alt Aeq.
     
     (apply interp.o_Z)||
     (cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;auto with zarith).
     intros ;apply P_id__dot__monotonic;assumption.
     intros ;apply P_id__dot__bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id__dot__hat_1_monotonic;assumption.
   Qed.
   
   End S.
End InterpZ.

Module WF_R_xml_0_deep_rew.
 Inductive DP_R_xml_0  :
  algebra.Alg.term ->algebra.Alg.term ->Prop := 
    (* <.(.(x_,y_),z_),.(x_,.(y_,z_))> *)
   | DP_R_xml_0_0 :
    forall x2 x6 x1 x5 x3, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id__dot_ (x1::x2::nil)) x6) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x3 x5) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id__dot_ (x1::
                   (algebra.Alg.Term algebra.F.id__dot_ (x2::x3::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id__dot_ (x6::x5::nil))
    (* <.(.(x_,y_),z_),.(y_,z_)> *)
   | DP_R_xml_0_1 :
    forall x2 x6 x1 x5 x3, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id__dot_ (x1::x2::nil)) x6) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x3 x5) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id__dot_ (x2::x3::nil)) 
        (algebra.Alg.Term algebra.F.id__dot_ (x6::x5::nil))
 .
 
 Module ddp := dp.MakeDP(algebra.EQT).
 
 
 Lemma R_xml_0_dp_step_spec :
  forall x y, 
   (ddp.dp_step R_xml_0_deep_rew.R_xml_0_rules x y) ->
    exists f,
      exists l1,
        exists l2,
          y = algebra.Alg.Term f l2/\ 
          (closure.refl_trans_clos (closure.one_step_list (algebra.EQT.one_step 
                                                            R_xml_0_deep_rew.R_xml_0_rules)
                                                           ) l1 l2)/\ 
          (ddp.dp R_xml_0_deep_rew.R_xml_0_rules x (algebra.Alg.Term f l1)).
 Proof.
   intros x y H.
   induction H.
   inversion H.
   subst.
   destruct t0.
   refine ((False_ind) _ _).
   refine (R_xml_0_deep_rew.R_xml_0_non_var H0).
   simpl in H|-*.
   exists a.
   exists ((List.map) (algebra.Alg.apply_subst sigma) l).
   exists ((List.map) (algebra.Alg.apply_subst sigma) l).
   repeat (constructor).
   assumption.
   exists f.
   exists l2.
   exists l1.
   constructor.
   constructor.
   constructor.
   constructor.
   rewrite  <- closure.rwr_list_trans_clos_one_step_list.
   assumption.
   assumption.
 Qed.
 
 
 Ltac included_dp_tac H :=
  injection H;clear H;intros;subst;
  repeat (match goal with 
  | H: closure.refl_trans_clos (closure.one_step_list _) (_::_) _ |- _=>           
  let x := fresh "x" in 
  let l := fresh "l" in 
  let h1 := fresh "h" in 
  let h2 := fresh "h" in 
  let h3 := fresh "h" in 
  destruct (@algebra.EQT_ext.one_step_list_star_decompose_cons _ _ _ _  H) as [x [l[h1[h2 h3]]]];clear H;subst
  | H: closure.refl_trans_clos (closure.one_step_list _) nil _ |- _ => 
  rewrite (@algebra.EQT_ext.one_step_list_star_decompose_nil _ _ H) in *;clear H
  end
  );simpl;
  econstructor eassumption
 .
 
 
 Ltac dp_concl_tac h2 h cont_tac 
  t :=
  match t with
    | False => let h' := fresh "a" in 
                (set (h':=t) in *;cont_tac h';
                  repeat (
                  let e := type of h in 
                   (match e with
                      | ?t => unfold t in h|-;
                               (case h;
                                [abstract (clear h;intros h;injection h;
                                            clear h;intros ;subst;
                                            included_dp_tac h2)|
                                clear h;intros h;clear t])
                      | ?t => unfold t in h|-;elim h
                      end
                    )
                  ))
    | or ?a ?b => let cont_tac 
                   h' := let h'' := fresh "a" in 
                          (set (h'':=or a h') in *;cont_tac h'') in 
                   (dp_concl_tac h2 h cont_tac b)
    end
  .
 
 
 Module WF_DP_R_xml_0.
  Inductive DP_R_xml_0_large  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <.(.(x_,y_),z_),.(x_,.(y_,z_))> *)
    | DP_R_xml_0_large_0 :
     forall x2 x6 x1 x5 x3, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id__dot_ (x1::x2::nil)) x6) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x3 x5) ->
        DP_R_xml_0_large (algebra.Alg.Term algebra.F.id__dot_ (x1::
                          (algebra.Alg.Term algebra.F.id__dot_ (x2::
                          x3::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id__dot_ (x6::x5::nil))
  .
  
  
  Inductive DP_R_xml_0_strict  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <.(.(x_,y_),z_),.(y_,z_)> *)
    | DP_R_xml_0_strict_0 :
     forall x2 x6 x1 x5 x3, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id__dot_ (x1::x2::nil)) x6) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x3 x5) ->
        DP_R_xml_0_strict (algebra.Alg.Term algebra.F.id__dot_ (x2::x3::nil)) 
         (algebra.Alg.Term algebra.F.id__dot_ (x6::x5::nil))
  .
  
  
  Module WF_DP_R_xml_0_large.
   Open Scope Z_scope.
   
   Import ring_extention.
   
   Notation Local "a <= b" := (Zle a b).
   
   Notation Local "a < b" := (Zlt a b).
   
   Definition P_id__dot_ (x5:Z) (x6:Z) := 2 + 2* x5.
   
   Lemma P_id__dot__monotonic :
    forall x8 x6 x5 x7, 
     (0 <= x8)/\ (x8 <= x7) ->
      (0 <= x6)/\ (x6 <= x5) ->P_id__dot_ x6 x8 <= P_id__dot_ x5 x7.
   Proof.
     intros x8 x7 x6 x5.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id__dot__bounded :
    forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__dot_ x6 x5.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition measure  := InterpZ.measure 0 P_id__dot_.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id__dot_ (x6::x5::nil)) =>
                    P_id__dot_ (measure x6) (measure x5)
                   | _ => 0
                   end.
   Proof.
     intros t;case t;intros ;apply refl_equal.
   Qed.
   
   Lemma measure_bounded : forall t, 0 <= measure t.
   Proof.
     unfold measure in |-*.
     
     apply InterpZ.measure_bounded;
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Ltac generate_pos_hyp  :=
    match goal with
      | H:context [measure ?x] |- _ =>
       let v := fresh "v" in 
        (let H := fresh "h" in 
          (set (H:=measure_bounded x) in *;set (v:=measure x) in *;
            clearbody H;clearbody v))
      |  |- context [measure ?x] =>
       let v := fresh "v" in 
        (let H := fresh "h" in 
          (set (H:=measure_bounded x) in *;set (v:=measure x) in *;
            clearbody H;clearbody v))
      end
    .
   
   Lemma rules_monotonic :
    forall l r, 
     (algebra.EQT.axiom R_xml_0_deep_rew.R_xml_0_rules r l) ->
      measure r <= measure l.
   Proof.
     intros l r H.
     fold measure in |-*.
     
     inversion H;clear H;subst;inversion H0;clear H0;subst;
      simpl algebra.EQT.T.apply_subst in |-*;
      repeat (
      match goal with
        |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
         rewrite (measure_equation (algebra.Alg.Term f t))
        end
      );repeat (generate_pos_hyp );
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma measure_star_monotonic :
    forall l r, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                r l) ->measure r <= measure l.
   Proof.
     unfold measure in *.
     apply InterpZ.measure_star_monotonic.
     intros ;apply P_id__dot__monotonic;assumption.
     intros ;apply P_id__dot__bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id__dot__hat_1 (x5:Z) (x6:Z) := 1* x5.
   
   Lemma P_id__dot__hat_1_monotonic :
    forall x8 x6 x5 x7, 
     (0 <= x8)/\ (x8 <= x7) ->
      (0 <= x6)/\ (x6 <= x5) ->
       P_id__dot__hat_1 x6 x8 <= P_id__dot__hat_1 x5 x7.
   Proof.
     intros x8 x7 x6 x5.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id__dot_ P_id__dot__hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id__dot_ (x6::
                             x5::nil)) =>
                           P_id__dot__hat_1 (measure x6) (measure x5)
                          | _ => measure t
                          end.
   Proof.
     reflexivity .
   Qed.
   
   Lemma marked_measure_star_monotonic :
    forall f l1 l2, 
     (closure.refl_trans_clos (closure.one_step_list (algebra.EQT.one_step 
                                                       R_xml_0_deep_rew.R_xml_0_rules)
                                                      ) l1 l2) ->
      marked_measure (algebra.Alg.Term f l1) <= marked_measure (algebra.Alg.Term 
                                                                 f l2).
   Proof.
     unfold marked_measure in *.
     apply InterpZ.marked_measure_star_monotonic.
     intros ;apply P_id__dot__monotonic;assumption.
     intros ;apply P_id__dot__bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id__dot__hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_large.
   Proof.
     intros x.
     
     apply well_founded_ind with
       (R:=fun x y => (Zwf.Zwf 0) (marked_measure x) (marked_measure y)).
     apply Inverse_Image.wf_inverse_image with  (B:=Z).
     apply Zwf.Zwf_well_founded.
     clear x.
     intros x IHx.
     
     repeat (
     constructor;inversion 1;subst;
      full_prove_ineq algebra.Alg.Term 
      ltac:(algebra.Alg_ext.find_replacement ) 
      algebra.EQT_ext.one_step_list_refl_trans_clos marked_measure 
      marked_measure_star_monotonic (Zwf.Zwf 0) (interp.o_Z 0) 
      ltac:(fun _ => R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 ) 
      ltac:(fun _ => rewrite_and_unfold ) ltac:(fun _ => generate_pos_hyp ) 
      ltac:(fun _ => cbv beta iota zeta delta - [Zplus Zmult Zle Zlt] in * ;
                      try (constructor))
       IHx
     ).
   Qed.
  End WF_DP_R_xml_0_large.
  
  Open Scope Z_scope.
  
  Import ring_extention.
  
  Notation Local "a <= b" := (Zle a b).
  
  Notation Local "a < b" := (Zlt a b).
  
  Definition P_id__dot_ (x5:Z) (x6:Z) := 1 + 1* x5 + 1* x6.
  
  Lemma P_id__dot__monotonic :
   forall x8 x6 x5 x7, 
    (0 <= x8)/\ (x8 <= x7) ->
     (0 <= x6)/\ (x6 <= x5) ->P_id__dot_ x6 x8 <= P_id__dot_ x5 x7.
  Proof.
    intros x8 x7 x6 x5.
    intros [H_1 H_0].
    intros [H_3 H_2].
    
    cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
     (auto with zarith)||(repeat (translate_vars );prove_ineq ).
  Qed.
  
  Lemma P_id__dot__bounded :
   forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__dot_ x6 x5.
  Proof.
    intros .
    
    cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
     (auto with zarith)||(repeat (translate_vars );prove_ineq ).
  Qed.
  
  Definition measure  := InterpZ.measure 0 P_id__dot_.
  
  Lemma measure_equation :
   forall t, 
    measure t = match t with
                  | (algebra.Alg.Term algebra.F.id__dot_ (x6::x5::nil)) =>
                   P_id__dot_ (measure x6) (measure x5)
                  | _ => 0
                  end.
  Proof.
    intros t;case t;intros ;apply refl_equal.
  Qed.
  
  Lemma measure_bounded : forall t, 0 <= measure t.
  Proof.
    unfold measure in |-*.
    
    apply InterpZ.measure_bounded;
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
  Qed.
  
  Ltac generate_pos_hyp  :=
   match goal with
     | H:context [measure ?x] |- _ =>
      let v := fresh "v" in 
       (let H := fresh "h" in 
         (set (H:=measure_bounded x) in *;set (v:=measure x) in *;
           clearbody H;clearbody v))
     |  |- context [measure ?x] =>
      let v := fresh "v" in 
       (let H := fresh "h" in 
         (set (H:=measure_bounded x) in *;set (v:=measure x) in *;
           clearbody H;clearbody v))
     end
   .
  
  Lemma rules_monotonic :
   forall l r, 
    (algebra.EQT.axiom R_xml_0_deep_rew.R_xml_0_rules r l) ->
     measure r <= measure l.
  Proof.
    intros l r H.
    fold measure in |-*.
    
    inversion H;clear H;subst;inversion H0;clear H0;subst;
     simpl algebra.EQT.T.apply_subst in |-*;
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       end
     );repeat (generate_pos_hyp );
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
  Qed.
  
  Lemma measure_star_monotonic :
   forall l r, 
    (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                               r l) ->measure r <= measure l.
  Proof.
    unfold measure in *.
    apply InterpZ.measure_star_monotonic.
    intros ;apply P_id__dot__monotonic;assumption.
    intros ;apply P_id__dot__bounded;assumption.
    apply rules_monotonic.
  Qed.
  
  Definition P_id__dot__hat_1 (x5:Z) (x6:Z) := 1* x5 + 1* x6.
  
  Lemma P_id__dot__hat_1_monotonic :
   forall x8 x6 x5 x7, 
    (0 <= x8)/\ (x8 <= x7) ->
     (0 <= x6)/\ (x6 <= x5) ->
      P_id__dot__hat_1 x6 x8 <= P_id__dot__hat_1 x5 x7.
  Proof.
    intros x8 x7 x6 x5.
    intros [H_1 H_0].
    intros [H_3 H_2].
    
    cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
     (auto with zarith)||(repeat (translate_vars );prove_ineq ).
  Qed.
  
  Definition marked_measure  := 
    InterpZ.marked_measure 0 P_id__dot_ P_id__dot__hat_1.
  
  Lemma marked_measure_equation :
   forall t, 
    marked_measure t = match t with
                         | (algebra.Alg.Term algebra.F.id__dot_ (x6::
                            x5::nil)) =>
                          P_id__dot__hat_1 (measure x6) (measure x5)
                         | _ => measure t
                         end.
  Proof.
    reflexivity .
  Qed.
  
  Lemma marked_measure_star_monotonic :
   forall f l1 l2, 
    (closure.refl_trans_clos (closure.one_step_list (algebra.EQT.one_step 
                                                      R_xml_0_deep_rew.R_xml_0_rules)
                                                     ) l1 l2) ->
     marked_measure (algebra.Alg.Term f l1) <= marked_measure (algebra.Alg.Term 
                                                                f l2).
  Proof.
    unfold marked_measure in *.
    apply InterpZ.marked_measure_star_monotonic.
    intros ;apply P_id__dot__monotonic;assumption.
    intros ;apply P_id__dot__bounded;assumption.
    apply rules_monotonic.
    intros ;apply P_id__dot__hat_1_monotonic;assumption.
  Qed.
  
  Ltac rewrite_and_unfold  :=
   do 2 (rewrite marked_measure_equation);
    repeat (
    match goal with
      |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
       rewrite (measure_equation (algebra.Alg.Term f t))
      | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
       rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
      end
    ).
  
  Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
  
  Definition le a b := marked_measure a <= marked_measure b.
  
  Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
  Proof.
    unfold lt, le in *.
    intros a b c.
    apply (interp.le_lt_compat_right (interp.o_Z 0)).
  Qed.
  
  Lemma wf_lt : well_founded lt.
  Proof.
    unfold lt in *.
    apply Inverse_Image.wf_inverse_image with  (B:=Z).
    apply Zwf.Zwf_well_founded.
  Qed.
  
  Lemma DP_R_xml_0_strict_in_lt :
   Relation_Definitions.inclusion _ DP_R_xml_0_strict lt.
  Proof.
    unfold Relation_Definitions.inclusion, lt in *.
    
    intros a b H;destruct H;
     match goal with
       |  |- (Zwf.Zwf 0) _ (marked_measure (algebra.Alg.Term ?f ?l)) =>
        let l'' := algebra.Alg_ext.find_replacement l  in 
         ((apply (interp.le_lt_compat_right (interp.o_Z 0)) with
            (marked_measure (algebra.Alg.Term f l''));[idtac|
           apply marked_measure_star_monotonic;
            repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos);
            (assumption)||(constructor 1)]))
       end
     ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp );
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
  Qed.
  
  Lemma DP_R_xml_0_large_in_le :
   Relation_Definitions.inclusion _ DP_R_xml_0_large le.
  Proof.
    unfold Relation_Definitions.inclusion, le, Zwf.Zwf in *.
    
    intros a b H;destruct H;
     match goal with
       |  |- _ <= marked_measure (algebra.Alg.Term ?f ?l) =>
        let l'' := algebra.Alg_ext.find_replacement l  in 
         ((apply (interp.le_trans (interp.o_Z 0)) with
            (marked_measure (algebra.Alg.Term f l''));[idtac|
           apply marked_measure_star_monotonic;
            repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos);
            (assumption)||(constructor 1)]))
       end
     ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp );
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
  Qed.
  
  Definition wf_DP_R_xml_0_large  := WF_DP_R_xml_0_large.wf.
  
  
  Lemma wf : well_founded WF_R_xml_0_deep_rew.DP_R_xml_0.
  Proof.
    intros x.
    apply (well_founded_ind wf_lt).
    clear x.
    intros x.
    pattern x.
    apply (@Acc_ind _ DP_R_xml_0_large).
    clear x.
    intros x _ IHx IHx'.
    constructor.
    intros y H.
    
    destruct H;
     (apply IHx';apply DP_R_xml_0_strict_in_lt;econstructor eassumption)||
     ((apply IHx;[econstructor eassumption|
       intros y' Hlt;apply IHx';apply lt_le_compat with (1:=Hlt) ;
        apply DP_R_xml_0_large_in_le;econstructor eassumption])).
    apply wf_DP_R_xml_0_large.
  Qed.
 End WF_DP_R_xml_0.
 
 Definition wf_H  := WF_DP_R_xml_0.wf.
 
 Lemma wf :
  well_founded (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules).
 Proof.
   apply ddp.dp_criterion.
   apply R_xml_0_deep_rew.R_xml_0_non_var.
   apply R_xml_0_deep_rew.R_xml_0_reg.
   
   intros ;
    apply (ddp.constructor_defined_dec _ _ 
            R_xml_0_deep_rew.R_xml_0_rules_included).
   refine (Inclusion.wf_incl _ _ _ _ wf_H).
   intros x y H.
   destruct (R_xml_0_dp_step_spec H) as [f [l1 [l2 [H1 [H2 H3]]]]].
   
   destruct (ddp.dp_list_complete _ _ 
              R_xml_0_deep_rew.R_xml_0_rules_included _ _ H3)
    as [x' [y' [sigma [h1 [h2 h3]]]]].
   clear H3.
   subst.
   vm_compute in h3|-.
   let e := type of h3 in (dp_concl_tac h2 h3 ltac:(fun _ => idtac) e).
 Qed.
End WF_R_xml_0_deep_rew.


(* 
*** Local Variables: ***
*** coq-prog-name: "coqtop" ***
*** coq-prog-args: ("-emacs-U" "-I" "$COCCINELLE/examples" "-I" "$COCCINELLE/term_algebra" "-I" "$COCCINELLE/term_orderings" "-I" "$COCCINELLE/basis" "-I" "$COCCINELLE/list_extensions" "-I" "$COCCINELLE/examples/cime_trace/") ***
*** compile-command: "coqc -I $COCCINELLE/term_algebra -I $COCCINELLE/term_orderings -I $COCCINELLE/basis -I $COCCINELLE/list_extensions -I $COCCINELLE/examples/cime_trace/ -I $COCCINELLE/examples/  c_output/strat/tpdb-5.0___TRS___D33___09.trs/a3pat.v" ***
*** End: ***
 *)