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_din *)
    | id_din : symb
     (* id_u32 *)
    | id_u32 : symb
     (* id_dout *)
    | id_dout : symb
     (* id_plus *)
    | id_plus : symb
     (* id_u42 *)
    | id_u42 : symb
     (* id_times *)
    | id_times : symb
     (* id_der *)
    | id_der : symb
     (* id_u41 *)
    | id_u41 : symb
     (* id_u22 *)
    | id_u22 : symb
     (* id_u21 *)
    | id_u21 : symb
     (* id_u31 *)
    | id_u31 : symb
  .
  
  
  Definition symb_eq_bool (f1 f2:symb) : bool := 
    match f1,f2 with
      | id_din,id_din => true
      | id_u32,id_u32 => true
      | id_dout,id_dout => true
      | id_plus,id_plus => true
      | id_u42,id_u42 => true
      | id_times,id_times => true
      | id_der,id_der => true
      | id_u41,id_u41 => true
      | id_u22,id_u22 => true
      | id_u21,id_u21 => true
      | id_u31,id_u31 => true
      | _,_ => false
      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_din,id_din => refl_equal _
             | id_u32,id_u32 => refl_equal _
             | id_dout,id_dout => refl_equal _
             | id_plus,id_plus => refl_equal _
             | id_u42,id_u42 => refl_equal _
             | id_times,id_times => refl_equal _
             | id_der,id_der => refl_equal _
             | id_u41,id_u41 => refl_equal _
             | id_u22,id_u22 => refl_equal _
             | id_u21,id_u21 => refl_equal _
             | id_u31,id_u31 => refl_equal _
             | _,_ => _
             end;intros abs;discriminate.
  Defined.
  
  
  Definition arity (f:symb) := 
    match f with
      | id_din => term.Free 1
      | id_u32 => term.Free 4
      | id_dout => term.Free 1
      | id_plus => term.Free 2
      | id_u42 => term.Free 3
      | id_times => term.Free 2
      | id_der => term.Free 1
      | id_u41 => term.Free 2
      | id_u22 => term.Free 4
      | id_u21 => term.Free 3
      | id_u31 => term.Free 3
      end.
  
  
  Definition symb_order (f1 f2:symb) : bool := 
    match f1,f2 with
      | id_din,id_din => true
      | id_din,id_u32 => false
      | id_din,id_dout => false
      | id_din,id_plus => false
      | id_din,id_u42 => false
      | id_din,id_times => false
      | id_din,id_der => false
      | id_din,id_u41 => false
      | id_din,id_u22 => false
      | id_din,id_u21 => false
      | id_din,id_u31 => false
      | id_u32,id_din => true
      | id_u32,id_u32 => true
      | id_u32,id_dout => false
      | id_u32,id_plus => false
      | id_u32,id_u42 => false
      | id_u32,id_times => false
      | id_u32,id_der => false
      | id_u32,id_u41 => false
      | id_u32,id_u22 => false
      | id_u32,id_u21 => false
      | id_u32,id_u31 => false
      | id_dout,id_din => true
      | id_dout,id_u32 => true
      | id_dout,id_dout => true
      | id_dout,id_plus => false
      | id_dout,id_u42 => false
      | id_dout,id_times => false
      | id_dout,id_der => false
      | id_dout,id_u41 => false
      | id_dout,id_u22 => false
      | id_dout,id_u21 => false
      | id_dout,id_u31 => false
      | id_plus,id_din => true
      | id_plus,id_u32 => true
      | id_plus,id_dout => true
      | id_plus,id_plus => true
      | id_plus,id_u42 => false
      | id_plus,id_times => false
      | id_plus,id_der => false
      | id_plus,id_u41 => false
      | id_plus,id_u22 => false
      | id_plus,id_u21 => false
      | id_plus,id_u31 => false
      | id_u42,id_din => true
      | id_u42,id_u32 => true
      | id_u42,id_dout => true
      | id_u42,id_plus => true
      | id_u42,id_u42 => true
      | id_u42,id_times => false
      | id_u42,id_der => false
      | id_u42,id_u41 => false
      | id_u42,id_u22 => false
      | id_u42,id_u21 => false
      | id_u42,id_u31 => false
      | id_times,id_din => true
      | id_times,id_u32 => true
      | id_times,id_dout => true
      | id_times,id_plus => true
      | id_times,id_u42 => true
      | id_times,id_times => true
      | id_times,id_der => false
      | id_times,id_u41 => false
      | id_times,id_u22 => false
      | id_times,id_u21 => false
      | id_times,id_u31 => false
      | id_der,id_din => true
      | id_der,id_u32 => true
      | id_der,id_dout => true
      | id_der,id_plus => true
      | id_der,id_u42 => true
      | id_der,id_times => true
      | id_der,id_der => true
      | id_der,id_u41 => false
      | id_der,id_u22 => false
      | id_der,id_u21 => false
      | id_der,id_u31 => false
      | id_u41,id_din => true
      | id_u41,id_u32 => true
      | id_u41,id_dout => true
      | id_u41,id_plus => true
      | id_u41,id_u42 => true
      | id_u41,id_times => true
      | id_u41,id_der => true
      | id_u41,id_u41 => true
      | id_u41,id_u22 => false
      | id_u41,id_u21 => false
      | id_u41,id_u31 => false
      | id_u22,id_din => true
      | id_u22,id_u32 => true
      | id_u22,id_dout => true
      | id_u22,id_plus => true
      | id_u22,id_u42 => true
      | id_u22,id_times => true
      | id_u22,id_der => true
      | id_u22,id_u41 => true
      | id_u22,id_u22 => true
      | id_u22,id_u21 => false
      | id_u22,id_u31 => false
      | id_u21,id_din => true
      | id_u21,id_u32 => true
      | id_u21,id_dout => true
      | id_u21,id_plus => true
      | id_u21,id_u42 => true
      | id_u21,id_times => true
      | id_u21,id_der => true
      | id_u21,id_u41 => true
      | id_u21,id_u22 => true
      | id_u21,id_u21 => true
      | id_u21,id_u31 => false
      | id_u31,id_din => true
      | id_u31,id_u32 => true
      | id_u31,id_dout => true
      | id_u31,id_plus => true
      | id_u31,id_u42 => true
      | id_u31,id_times => true
      | id_u31,id_der => true
      | id_u31,id_u41 => true
      | id_u31,id_u22 => true
      | id_u31,id_u21 => true
      | id_u31,id_u31 => 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 := 
    (* din(der(plus(X_,Y_))) -> u21(din(der(X_)),X_,Y_) *)
   | R_xml_0_rule_0 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_u21 ((algebra.Alg.Term 
                   algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
                   ((algebra.Alg.Var 1)::nil))::nil))::(algebra.Alg.Var 1)::
                   (algebra.Alg.Var 2)::nil)) 
     (algebra.Alg.Term algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
      ((algebra.Alg.Term algebra.F.id_plus ((algebra.Alg.Var 1)::
      (algebra.Alg.Var 2)::nil))::nil))::nil))
   
    (* u21(dout(DX_),X_,Y_) -> u22(din(der(Y_)),X_,Y_,DX_) *)
   | R_xml_0_rule_1 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_u22 ((algebra.Alg.Term 
                   algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
                   ((algebra.Alg.Var 2)::nil))::nil))::(algebra.Alg.Var 1)::
                   (algebra.Alg.Var 2)::(algebra.Alg.Var 3)::nil)) 
     (algebra.Alg.Term algebra.F.id_u21 ((algebra.Alg.Term algebra.F.id_dout 
      ((algebra.Alg.Var 3)::nil))::(algebra.Alg.Var 1)::
      (algebra.Alg.Var 2)::nil))
    (* u22(dout(DY_),X_,Y_,DX_) -> dout(plus(DX_,DY_)) *)
   | R_xml_0_rule_2 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_dout ((algebra.Alg.Term 
                   algebra.F.id_plus ((algebra.Alg.Var 3)::
                   (algebra.Alg.Var 4)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_u22 ((algebra.Alg.Term algebra.F.id_dout 
      ((algebra.Alg.Var 4)::nil))::(algebra.Alg.Var 1)::(algebra.Alg.Var 2)::
      (algebra.Alg.Var 3)::nil))
    (* din(der(times(X_,Y_))) -> u31(din(der(X_)),X_,Y_) *)
   | R_xml_0_rule_3 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_u31 ((algebra.Alg.Term 
                   algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
                   ((algebra.Alg.Var 1)::nil))::nil))::(algebra.Alg.Var 1)::
                   (algebra.Alg.Var 2)::nil)) 
     (algebra.Alg.Term algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
      ((algebra.Alg.Term algebra.F.id_times ((algebra.Alg.Var 1)::
      (algebra.Alg.Var 2)::nil))::nil))::nil))
   
    (* u31(dout(DX_),X_,Y_) -> u32(din(der(Y_)),X_,Y_,DX_) *)
   | R_xml_0_rule_4 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_u32 ((algebra.Alg.Term 
                   algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
                   ((algebra.Alg.Var 2)::nil))::nil))::(algebra.Alg.Var 1)::
                   (algebra.Alg.Var 2)::(algebra.Alg.Var 3)::nil)) 
     (algebra.Alg.Term algebra.F.id_u31 ((algebra.Alg.Term algebra.F.id_dout 
      ((algebra.Alg.Var 3)::nil))::(algebra.Alg.Var 1)::
      (algebra.Alg.Var 2)::nil))
   
    (* u32(dout(DY_),X_,Y_,DX_) -> dout(plus(times(X_,DY_),times(Y_,DX_))) *)
   | R_xml_0_rule_5 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_dout ((algebra.Alg.Term 
                   algebra.F.id_plus ((algebra.Alg.Term algebra.F.id_times 
                   ((algebra.Alg.Var 1)::(algebra.Alg.Var 4)::nil))::
                   (algebra.Alg.Term algebra.F.id_times 
                   ((algebra.Alg.Var 2)::
                   (algebra.Alg.Var 3)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_u32 ((algebra.Alg.Term algebra.F.id_dout 
      ((algebra.Alg.Var 4)::nil))::(algebra.Alg.Var 1)::(algebra.Alg.Var 2)::
      (algebra.Alg.Var 3)::nil))
    (* din(der(der(X_))) -> u41(din(der(X_)),X_) *)
   | R_xml_0_rule_6 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_u41 ((algebra.Alg.Term 
                   algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
                   ((algebra.Alg.Var 1)::nil))::nil))::
                   (algebra.Alg.Var 1)::nil)) 
     (algebra.Alg.Term algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
      ((algebra.Alg.Term algebra.F.id_der 
      ((algebra.Alg.Var 1)::nil))::nil))::nil))
    (* u41(dout(DX_),X_) -> u42(din(der(DX_)),X_,DX_) *)
   | R_xml_0_rule_7 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_u42 ((algebra.Alg.Term 
                   algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
                   ((algebra.Alg.Var 3)::nil))::nil))::(algebra.Alg.Var 1)::
                   (algebra.Alg.Var 3)::nil)) 
     (algebra.Alg.Term algebra.F.id_u41 ((algebra.Alg.Term algebra.F.id_dout 
      ((algebra.Alg.Var 3)::nil))::(algebra.Alg.Var 1)::nil))
    (* u42(dout(DDX_),X_,DX_) -> dout(DDX_) *)
   | R_xml_0_rule_8 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_dout 
                   ((algebra.Alg.Var 5)::nil)) 
     (algebra.Alg.Term algebra.F.id_u42 ((algebra.Alg.Term algebra.F.id_dout 
      ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Var 1)::
      (algebra.Alg.Var 3)::nil))
 .
 
 
 Definition R_xml_0_rule_as_list_0  := 
   ((algebra.Alg.Term algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
     ((algebra.Alg.Term algebra.F.id_plus ((algebra.Alg.Var 1)::
     (algebra.Alg.Var 2)::nil))::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_u21 ((algebra.Alg.Term algebra.F.id_din 
     ((algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Var 1)::nil))::nil))::
     (algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil)))::nil.
 
 
 Definition R_xml_0_rule_as_list_1  := 
   ((algebra.Alg.Term algebra.F.id_u21 ((algebra.Alg.Term algebra.F.id_dout 
     ((algebra.Alg.Var 3)::nil))::(algebra.Alg.Var 1)::
     (algebra.Alg.Var 2)::nil)),
    (algebra.Alg.Term algebra.F.id_u22 ((algebra.Alg.Term algebra.F.id_din 
     ((algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Var 2)::nil))::nil))::
     (algebra.Alg.Var 1)::(algebra.Alg.Var 2)::(algebra.Alg.Var 3)::nil)))::
    R_xml_0_rule_as_list_0.
 
 
 Definition R_xml_0_rule_as_list_2  := 
   ((algebra.Alg.Term algebra.F.id_u22 ((algebra.Alg.Term algebra.F.id_dout 
     ((algebra.Alg.Var 4)::nil))::(algebra.Alg.Var 1)::(algebra.Alg.Var 2)::
     (algebra.Alg.Var 3)::nil)),
    (algebra.Alg.Term algebra.F.id_dout ((algebra.Alg.Term algebra.F.id_plus 
     ((algebra.Alg.Var 3)::(algebra.Alg.Var 4)::nil))::nil)))::
    R_xml_0_rule_as_list_1.
 
 
 Definition R_xml_0_rule_as_list_3  := 
   ((algebra.Alg.Term algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
     ((algebra.Alg.Term algebra.F.id_times ((algebra.Alg.Var 1)::
     (algebra.Alg.Var 2)::nil))::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_u31 ((algebra.Alg.Term algebra.F.id_din 
     ((algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Var 1)::nil))::nil))::
     (algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil)))::R_xml_0_rule_as_list_2
   .
 
 
 Definition R_xml_0_rule_as_list_4  := 
   ((algebra.Alg.Term algebra.F.id_u31 ((algebra.Alg.Term algebra.F.id_dout 
     ((algebra.Alg.Var 3)::nil))::(algebra.Alg.Var 1)::
     (algebra.Alg.Var 2)::nil)),
    (algebra.Alg.Term algebra.F.id_u32 ((algebra.Alg.Term algebra.F.id_din 
     ((algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Var 2)::nil))::nil))::
     (algebra.Alg.Var 1)::(algebra.Alg.Var 2)::(algebra.Alg.Var 3)::nil)))::
    R_xml_0_rule_as_list_3.
 
 
 Definition R_xml_0_rule_as_list_5  := 
   ((algebra.Alg.Term algebra.F.id_u32 ((algebra.Alg.Term algebra.F.id_dout 
     ((algebra.Alg.Var 4)::nil))::(algebra.Alg.Var 1)::(algebra.Alg.Var 2)::
     (algebra.Alg.Var 3)::nil)),
    (algebra.Alg.Term algebra.F.id_dout ((algebra.Alg.Term algebra.F.id_plus 
     ((algebra.Alg.Term algebra.F.id_times ((algebra.Alg.Var 1)::
     (algebra.Alg.Var 4)::nil))::(algebra.Alg.Term algebra.F.id_times 
     ((algebra.Alg.Var 2)::(algebra.Alg.Var 3)::nil))::nil))::nil)))::
    R_xml_0_rule_as_list_4.
 
 
 Definition R_xml_0_rule_as_list_6  := 
   ((algebra.Alg.Term algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
     ((algebra.Alg.Term algebra.F.id_der 
     ((algebra.Alg.Var 1)::nil))::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_u41 ((algebra.Alg.Term algebra.F.id_din 
     ((algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Var 1)::nil))::nil))::
     (algebra.Alg.Var 1)::nil)))::R_xml_0_rule_as_list_5.
 
 
 Definition R_xml_0_rule_as_list_7  := 
   ((algebra.Alg.Term algebra.F.id_u41 ((algebra.Alg.Term algebra.F.id_dout 
     ((algebra.Alg.Var 3)::nil))::(algebra.Alg.Var 1)::nil)),
    (algebra.Alg.Term algebra.F.id_u42 ((algebra.Alg.Term algebra.F.id_din 
     ((algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Var 3)::nil))::nil))::
     (algebra.Alg.Var 1)::(algebra.Alg.Var 3)::nil)))::R_xml_0_rule_as_list_6
   .
 
 
 Definition R_xml_0_rule_as_list_8  := 
   ((algebra.Alg.Term algebra.F.id_u42 ((algebra.Alg.Term algebra.F.id_dout 
     ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Var 1)::
     (algebra.Alg.Var 3)::nil)),
    (algebra.Alg.Term algebra.F.id_dout ((algebra.Alg.Var 5)::nil)))::
    R_xml_0_rule_as_list_7.
 
 Definition R_xml_0_rule_as_list  := R_xml_0_rule_as_list_8.
 
 
 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.
   vm_compute in H|-.
   rewrite  <- or_ext_generated.or10_equiv in H|-.
   case H;clear H;intros H.
   injection H;intros ;subst;constructor 9.
   injection H;intros ;subst;constructor 8.
   injection H;intros ;subst;constructor 7.
   injection H;intros ;subst;constructor 6.
   injection H;intros ;subst;constructor 5.
   injection H;intros ;subst;constructor 4.
   injection H;intros ;subst;constructor 3.
   injection H;intros ;subst;constructor 2.
   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.
 
 
 Inductive and_4 (x7 x8 x9 x10:Prop) :
  Prop := 
   | conj_4 : x7->x8->x9->x10->and_4 x7 x8 x9 x10
 .
 
 
 Lemma are_constuctors_of_R_xml_0 :
  forall t t', 
   (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) ->
    and_4 (forall x8, 
           t = (algebra.Alg.Term algebra.F.id_dout (x8::nil)) ->
            exists x7,
              t' = (algebra.Alg.Term algebra.F.id_dout (x7::nil))/\ 
              (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
                x7 x8)) 
     (forall x8 x10, 
      t = (algebra.Alg.Term algebra.F.id_plus (x8::x10::nil)) ->
       exists x7,
         exists x9,
           t' = (algebra.Alg.Term algebra.F.id_plus (x7::x9::nil))/\ 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
             x7 x8)/\ 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
             x9 x10)) 
     (forall x8 x10, 
      t = (algebra.Alg.Term algebra.F.id_times (x8::x10::nil)) ->
       exists x7,
         exists x9,
           t' = (algebra.Alg.Term algebra.F.id_times (x7::x9::nil))/\ 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
             x7 x8)/\ 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
             x9 x10)) 
     (forall x8, 
      t = (algebra.Alg.Term algebra.F.id_der (x8::nil)) ->
       exists x7,
         t' = (algebra.Alg.Term algebra.F.id_der (x7::nil))/\ 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x7 x8)).
 Proof.
   intros t t' H.
   
   induction H as [|y IH z z_to_y] using 
   closure_extension.refl_trans_clos_ind2.
   constructor 1.
   intros x8 H;exists x8;intuition;constructor 1.
   intros x8 x10 H;exists x8;exists x10;intuition;constructor 1.
   intros x8 x10 H;exists x8;exists x10;intuition;constructor 1.
   intros x8 H;exists x8;intuition;constructor 1.
   inversion z_to_y as [t1 t2 H H0 H1|f l1 l2 H0 H H2];clear z_to_y;subst.
   
   inversion H as [t1 t2 sigma H2 H1 H0];clear H IH;subst;inversion H2;
    clear ;constructor;try (intros until 0 );clear ;intros abs;
    discriminate abs.
   destruct IH as [H_id_dout H_id_plus H_id_times H_id_der].
   constructor.
   
   clear H_id_plus H_id_times H_id_der;intros x8 H;injection H;clear H;
    intros ;subst;
    repeat (
    match goal with
      | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ =>
       inversion H;clear H;subst
      end
    ).
   
   match goal with
     | H:algebra.EQT.one_step _ ?y x8 |- _ =>
      destruct (H_id_dout y (refl_equal _)) as [x7];intros ;intuition;
       exists x7;intuition;eapply closure_extension.refl_trans_clos_R;
       eassumption
     end
   .
   
   clear H_id_dout H_id_times H_id_der;intros x8 x10 H;injection H;clear H;
    intros ;subst;
    repeat (
    match goal with
      | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ =>
       inversion H;clear H;subst
      end
    ).
   
   match goal with
     | H:algebra.EQT.one_step _ ?y x8 |- _ =>
      destruct (H_id_plus y x10 (refl_equal _)) as [x7 [x9]];intros ;
       intuition;exists x7;exists x9;intuition;
       eapply closure_extension.refl_trans_clos_R;eassumption
     end
   .
   
   match goal with
     | H:algebra.EQT.one_step _ ?y x10 |- _ =>
      destruct (H_id_plus x8 y (refl_equal _)) as [x7 [x9]];intros ;
       intuition;exists x7;exists x9;intuition;
       eapply closure_extension.refl_trans_clos_R;eassumption
     end
   .
   
   clear H_id_dout H_id_plus H_id_der;intros x8 x10 H;injection H;clear H;
    intros ;subst;
    repeat (
    match goal with
      | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ =>
       inversion H;clear H;subst
      end
    ).
   
   match goal with
     | H:algebra.EQT.one_step _ ?y x8 |- _ =>
      destruct (H_id_times y x10 (refl_equal _)) as [x7 [x9]];intros ;
       intuition;exists x7;exists x9;intuition;
       eapply closure_extension.refl_trans_clos_R;eassumption
     end
   .
   
   match goal with
     | H:algebra.EQT.one_step _ ?y x10 |- _ =>
      destruct (H_id_times x8 y (refl_equal _)) as [x7 [x9]];intros ;
       intuition;exists x7;exists x9;intuition;
       eapply closure_extension.refl_trans_clos_R;eassumption
     end
   .
   
   clear H_id_dout H_id_plus H_id_times;intros x8 H;injection H;clear H;
    intros ;subst;
    repeat (
    match goal with
      | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ =>
       inversion H;clear H;subst
      end
    ).
   
   match goal with
     | H:algebra.EQT.one_step _ ?y x8 |- _ =>
      destruct (H_id_der y (refl_equal _)) as [x7];intros ;intuition;
       exists x7;intuition;eapply closure_extension.refl_trans_clos_R;
       eassumption
     end
   .
 Qed.
 
 
 Lemma id_dout_is_R_xml_0_constructor :
  forall t t', 
   (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) ->
    forall x8, 
     t = (algebra.Alg.Term algebra.F.id_dout (x8::nil)) ->
      exists x7,
        t' = (algebra.Alg.Term algebra.F.id_dout (x7::nil))/\ 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x7 x8).
 Proof.
   intros t t' H.
   destruct (are_constuctors_of_R_xml_0 H).
   assumption.
 Qed.
 
 
 Lemma id_plus_is_R_xml_0_constructor :
  forall t t', 
   (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) ->
    forall x8 x10, 
     t = (algebra.Alg.Term algebra.F.id_plus (x8::x10::nil)) ->
      exists x7,
        exists x9,
          t' = (algebra.Alg.Term algebra.F.id_plus (x7::x9::nil))/\ 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x7 x8)/\
           
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
            x9 x10).
 Proof.
   intros t t' H.
   destruct (are_constuctors_of_R_xml_0 H).
   assumption.
 Qed.
 
 
 Lemma id_times_is_R_xml_0_constructor :
  forall t t', 
   (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) ->
    forall x8 x10, 
     t = (algebra.Alg.Term algebra.F.id_times (x8::x10::nil)) ->
      exists x7,
        exists x9,
          t' = (algebra.Alg.Term algebra.F.id_times (x7::x9::nil))/\ 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x7 x8)/\
           
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
            x9 x10).
 Proof.
   intros t t' H.
   destruct (are_constuctors_of_R_xml_0 H).
   assumption.
 Qed.
 
 
 Lemma id_der_is_R_xml_0_constructor :
  forall t t', 
   (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) ->
    forall x8, 
     t = (algebra.Alg.Term algebra.F.id_der (x8::nil)) ->
      exists x7,
        t' = (algebra.Alg.Term algebra.F.id_der (x7::nil))/\ 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x7 x8).
 Proof.
   intros t t' H.
   destruct (are_constuctors_of_R_xml_0 H).
   assumption.
 Qed.
 
 
 Ltac impossible_star_reduction_R_xml_0  :=
  match goal with
    | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
         _ (algebra.Alg.Term algebra.F.id_dout (?x7::nil)) |- _ =>
     let x7 := fresh "x" in 
      (let Heq := fresh "Heq" in 
        (let Hred1 := fresh "Hred" in 
          (destruct (id_dout_is_R_xml_0_constructor H (refl_equal _)) as 
           [x7 [Heq Hred1]];
            (discriminate Heq)||
            (injection Heq;intros ;subst;clear Heq;clear H;
              impossible_star_reduction_R_xml_0 ))))
    | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
         _ (algebra.Alg.Term algebra.F.id_plus (?x8::?x7::nil)) |- _ =>
     let x8 := fresh "x" in 
      (let x7 := fresh "x" in 
        (let Heq := fresh "Heq" in 
          (let Hred1 := fresh "Hred" in 
            (let Hred2 := fresh "Hred" in 
              (destruct (id_plus_is_R_xml_0_constructor H (refl_equal _)) as 
               [x8 [x7 [Heq [Hred2 Hred1]]]];
                (discriminate Heq)||
                (injection Heq;intros ;subst;clear Heq;clear H;
                  impossible_star_reduction_R_xml_0 ))))))
    | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
         _ (algebra.Alg.Term algebra.F.id_times (?x8::?x7::nil)) |- _ =>
     let x8 := fresh "x" in 
      (let x7 := fresh "x" in 
        (let Heq := fresh "Heq" in 
          (let Hred1 := fresh "Hred" in 
            (let Hred2 := fresh "Hred" in 
              (destruct (id_times_is_R_xml_0_constructor H (refl_equal _))
                as [x8 [x7 [Heq [Hred2 Hred1]]]];
                (discriminate Heq)||
                (injection Heq;intros ;subst;clear Heq;clear H;
                  impossible_star_reduction_R_xml_0 ))))))
    | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
         _ (algebra.Alg.Term algebra.F.id_der (?x7::nil)) |- _ =>
     let x7 := fresh "x" in 
      (let Heq := fresh "Heq" in 
        (let Hred1 := fresh "Hred" in 
          (destruct (id_der_is_R_xml_0_constructor H (refl_equal _)) as 
           [x7 [Heq Hred1]];
            (discriminate Heq)||
            (injection Heq;intros ;subst;clear Heq;clear H;
              impossible_star_reduction_R_xml_0 ))))
    end
  .
 
 
 Ltac simplify_star_reduction_R_xml_0  :=
  match goal with
    | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
         _ (algebra.Alg.Term algebra.F.id_dout (?x7::nil)) |- _ =>
     let x7 := fresh "x" in 
      (let Heq := fresh "Heq" in 
        (let Hred1 := fresh "Hred" in 
          (destruct (id_dout_is_R_xml_0_constructor H (refl_equal _)) as 
           [x7 [Heq Hred1]];
            (discriminate Heq)||
            (injection Heq;intros ;subst;clear Heq;clear H;
              try (simplify_star_reduction_R_xml_0 )))))
    | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
         _ (algebra.Alg.Term algebra.F.id_plus (?x8::?x7::nil)) |- _ =>
     let x8 := fresh "x" in 
      (let x7 := fresh "x" in 
        (let Heq := fresh "Heq" in 
          (let Hred1 := fresh "Hred" in 
            (let Hred2 := fresh "Hred" in 
              (destruct (id_plus_is_R_xml_0_constructor H (refl_equal _)) as 
               [x8 [x7 [Heq [Hred2 Hred1]]]];
                (discriminate Heq)||
                (injection Heq;intros ;subst;clear Heq;clear H;
                  try (simplify_star_reduction_R_xml_0 )))))))
    | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
         _ (algebra.Alg.Term algebra.F.id_times (?x8::?x7::nil)) |- _ =>
     let x8 := fresh "x" in 
      (let x7 := fresh "x" in 
        (let Heq := fresh "Heq" in 
          (let Hred1 := fresh "Hred" in 
            (let Hred2 := fresh "Hred" in 
              (destruct (id_times_is_R_xml_0_constructor H (refl_equal _))
                as [x8 [x7 [Heq [Hred2 Hred1]]]];
                (discriminate Heq)||
                (injection Heq;intros ;subst;clear Heq;clear H;
                  try (simplify_star_reduction_R_xml_0 )))))))
    | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
         _ (algebra.Alg.Term algebra.F.id_der (?x7::nil)) |- _ =>
     let x7 := fresh "x" in 
      (let Heq := fresh "Heq" in 
        (let Hred1 := fresh "Hred" in 
          (destruct (id_der_is_R_xml_0_constructor H (refl_equal _)) as 
           [x7 [Heq Hred1]];
            (discriminate Heq)||
            (injection Heq;intros ;subst;clear Heq;clear H;
              try (simplify_star_reduction_R_xml_0 )))))
    end
  .
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_din : A ->A.
   
   Hypothesis P_id_u32 : A ->A ->A ->A ->A.
   
   Hypothesis P_id_dout : A ->A.
   
   Hypothesis P_id_plus : A ->A ->A.
   
   Hypothesis P_id_u42 : A ->A ->A ->A.
   
   Hypothesis P_id_times : A ->A ->A.
   
   Hypothesis P_id_der : A ->A.
   
   Hypothesis P_id_u41 : A ->A ->A.
   
   Hypothesis P_id_u22 : A ->A ->A ->A ->A.
   
   Hypothesis P_id_u21 : A ->A ->A ->A.
   
   Hypothesis P_id_u31 : A ->A ->A ->A.
   
   Hypothesis P_id_din_monotonic :
    forall x8 x7, (A0 <= x8)/\ (x8 <= x7) ->P_id_din x8 <= P_id_din x7.
   
   Hypothesis P_id_u32_monotonic :
    forall x8 x12 x10 x14 x9 x13 x11 x7, 
     (A0 <= x14)/\ (x14 <= x13) ->
      (A0 <= x12)/\ (x12 <= x11) ->
       (A0 <= x10)/\ (x10 <= x9) ->
        (A0 <= x8)/\ (x8 <= x7) ->
         P_id_u32 x8 x10 x12 x14 <= P_id_u32 x7 x9 x11 x13.
   
   Hypothesis P_id_dout_monotonic :
    forall x8 x7, (A0 <= x8)/\ (x8 <= x7) ->P_id_dout x8 <= P_id_dout x7.
   
   Hypothesis P_id_plus_monotonic :
    forall x8 x10 x9 x7, 
     (A0 <= x10)/\ (x10 <= x9) ->
      (A0 <= x8)/\ (x8 <= x7) ->P_id_plus x8 x10 <= P_id_plus x7 x9.
   
   Hypothesis P_id_u42_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (A0 <= x12)/\ (x12 <= x11) ->
      (A0 <= x10)/\ (x10 <= x9) ->
       (A0 <= x8)/\ (x8 <= x7) ->P_id_u42 x8 x10 x12 <= P_id_u42 x7 x9 x11.
   
   Hypothesis P_id_times_monotonic :
    forall x8 x10 x9 x7, 
     (A0 <= x10)/\ (x10 <= x9) ->
      (A0 <= x8)/\ (x8 <= x7) ->P_id_times x8 x10 <= P_id_times x7 x9.
   
   Hypothesis P_id_der_monotonic :
    forall x8 x7, (A0 <= x8)/\ (x8 <= x7) ->P_id_der x8 <= P_id_der x7.
   
   Hypothesis P_id_u41_monotonic :
    forall x8 x10 x9 x7, 
     (A0 <= x10)/\ (x10 <= x9) ->
      (A0 <= x8)/\ (x8 <= x7) ->P_id_u41 x8 x10 <= P_id_u41 x7 x9.
   
   Hypothesis P_id_u22_monotonic :
    forall x8 x12 x10 x14 x9 x13 x11 x7, 
     (A0 <= x14)/\ (x14 <= x13) ->
      (A0 <= x12)/\ (x12 <= x11) ->
       (A0 <= x10)/\ (x10 <= x9) ->
        (A0 <= x8)/\ (x8 <= x7) ->
         P_id_u22 x8 x10 x12 x14 <= P_id_u22 x7 x9 x11 x13.
   
   Hypothesis P_id_u21_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (A0 <= x12)/\ (x12 <= x11) ->
      (A0 <= x10)/\ (x10 <= x9) ->
       (A0 <= x8)/\ (x8 <= x7) ->P_id_u21 x8 x10 x12 <= P_id_u21 x7 x9 x11.
   
   Hypothesis P_id_u31_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (A0 <= x12)/\ (x12 <= x11) ->
      (A0 <= x10)/\ (x10 <= x9) ->
       (A0 <= x8)/\ (x8 <= x7) ->P_id_u31 x8 x10 x12 <= P_id_u31 x7 x9 x11.
   
   Hypothesis P_id_din_bounded : forall x7, (A0 <= x7) ->A0 <= P_id_din x7.
   
   Hypothesis P_id_u32_bounded :
    forall x8 x10 x9 x7, 
     (A0 <= x7) ->
      (A0 <= x8) ->(A0 <= x9) ->(A0 <= x10) ->A0 <= P_id_u32 x10 x9 x8 x7.
   
   Hypothesis P_id_dout_bounded : forall x7, (A0 <= x7) ->A0 <= P_id_dout x7.
   
   Hypothesis P_id_plus_bounded :
    forall x8 x7, (A0 <= x7) ->(A0 <= x8) ->A0 <= P_id_plus x8 x7.
   
   Hypothesis P_id_u42_bounded :
    forall x8 x9 x7, 
     (A0 <= x7) ->(A0 <= x8) ->(A0 <= x9) ->A0 <= P_id_u42 x9 x8 x7.
   
   Hypothesis P_id_times_bounded :
    forall x8 x7, (A0 <= x7) ->(A0 <= x8) ->A0 <= P_id_times x8 x7.
   
   Hypothesis P_id_der_bounded : forall x7, (A0 <= x7) ->A0 <= P_id_der x7.
   
   Hypothesis P_id_u41_bounded :
    forall x8 x7, (A0 <= x7) ->(A0 <= x8) ->A0 <= P_id_u41 x8 x7.
   
   Hypothesis P_id_u22_bounded :
    forall x8 x10 x9 x7, 
     (A0 <= x7) ->
      (A0 <= x8) ->(A0 <= x9) ->(A0 <= x10) ->A0 <= P_id_u22 x10 x9 x8 x7.
   
   Hypothesis P_id_u21_bounded :
    forall x8 x9 x7, 
     (A0 <= x7) ->(A0 <= x8) ->(A0 <= x9) ->A0 <= P_id_u21 x9 x8 x7.
   
   Hypothesis P_id_u31_bounded :
    forall x8 x9 x7, 
     (A0 <= x7) ->(A0 <= x8) ->(A0 <= x9) ->A0 <= P_id_u31 x9 x8 x7.
   
   Fixpoint measure t { struct t }  := 
     match t with
       | (algebra.Alg.Term algebra.F.id_din (x7::nil)) =>
        P_id_din (measure x7)
       | (algebra.Alg.Term algebra.F.id_u32 (x10::x9::x8::x7::nil)) =>
        P_id_u32 (measure x10) (measure x9) (measure x8) (measure x7)
       | (algebra.Alg.Term algebra.F.id_dout (x7::nil)) =>
        P_id_dout (measure x7)
       | (algebra.Alg.Term algebra.F.id_plus (x8::x7::nil)) =>
        P_id_plus (measure x8) (measure x7)
       | (algebra.Alg.Term algebra.F.id_u42 (x9::x8::x7::nil)) =>
        P_id_u42 (measure x9) (measure x8) (measure x7)
       | (algebra.Alg.Term algebra.F.id_times (x8::x7::nil)) =>
        P_id_times (measure x8) (measure x7)
       | (algebra.Alg.Term algebra.F.id_der (x7::nil)) =>
        P_id_der (measure x7)
       | (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil)) =>
        P_id_u41 (measure x8) (measure x7)
       | (algebra.Alg.Term algebra.F.id_u22 (x10::x9::x8::x7::nil)) =>
        P_id_u22 (measure x10) (measure x9) (measure x8) (measure x7)
       | (algebra.Alg.Term algebra.F.id_u21 (x9::x8::x7::nil)) =>
        P_id_u21 (measure x9) (measure x8) (measure x7)
       | (algebra.Alg.Term algebra.F.id_u31 (x9::x8::x7::nil)) =>
        P_id_u31 (measure x9) (measure x8) (measure x7)
       | _ => A0
       end.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_din (x7::nil)) =>
                    P_id_din (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u32 (x10::x9::x8::
                      x7::nil)) =>
                    P_id_u32 (measure x10) (measure x9) (measure x8) 
                     (measure x7)
                   | (algebra.Alg.Term algebra.F.id_dout (x7::nil)) =>
                    P_id_dout (measure x7)
                   | (algebra.Alg.Term algebra.F.id_plus (x8::x7::nil)) =>
                    P_id_plus (measure x8) (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u42 (x9::x8::x7::nil)) =>
                    P_id_u42 (measure x9) (measure x8) (measure x7)
                   | (algebra.Alg.Term algebra.F.id_times (x8::x7::nil)) =>
                    P_id_times (measure x8) (measure x7)
                   | (algebra.Alg.Term algebra.F.id_der (x7::nil)) =>
                    P_id_der (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil)) =>
                    P_id_u41 (measure x8) (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u22 (x10::x9::x8::
                      x7::nil)) =>
                    P_id_u22 (measure x10) (measure x9) (measure x8) 
                     (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u21 (x9::x8::x7::nil)) =>
                    P_id_u21 (measure x9) (measure x8) (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u31 (x9::x8::x7::nil)) =>
                    P_id_u31 (measure x9) (measure x8) (measure x7)
                   | _ => 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_din => P_id_din
       | algebra.F.id_u32 => P_id_u32
       | algebra.F.id_dout => P_id_dout
       | algebra.F.id_plus => P_id_plus
       | algebra.F.id_u42 => P_id_u42
       | algebra.F.id_times => P_id_times
       | algebra.F.id_der => P_id_der
       | algebra.F.id_u41 => P_id_u41
       | algebra.F.id_u22 => P_id_u22
       | algebra.F.id_u21 => P_id_u21
       | algebra.F.id_u31 => P_id_u31
       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_din =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_u32 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::nil => _
                 | _::_::_::_::nil => _
                 | _::_::_::_::_::_ => _
                 end
              | algebra.F.id_dout =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_plus =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_u42 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::nil => _
                 | _::_::_::_::_ => _
                 end
              | algebra.F.id_times =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_der =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_u41 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_u22 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::nil => _
                 | _::_::_::_::nil => _
                 | _::_::_::_::_::_ => _
                 end
              | algebra.F.id_u21 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::nil => _
                 | _::_::_::_::_ => _
                 end
              | algebra.F.id_u31 =>
               match l with
                 | nil => _
                 | _::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_din_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u32_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_dout_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_plus_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u42_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_times_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_der_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u41_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u22_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u21_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u31_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_din_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_u32_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_dout_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_plus_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_u42_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_times_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_der_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_u41_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_u22_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_u21_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_u31_monotonic;assumption.
     intros f.
     case f.
     vm_compute in |-*;intros ;apply P_id_din_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u32_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_dout_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_plus_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u42_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_times_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_der_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u41_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u22_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u21_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u31_bounded;assumption.
     intros .
     do 2 (rewrite  <- same_measure in |-*).
     apply rules_monotonic;assumption.
     assumption.
   Qed.
   
   Hypothesis P_id_U41 : A ->A ->A.
   
   Hypothesis P_id_U21 : A ->A ->A ->A.
   
   Hypothesis P_id_U31 : A ->A ->A ->A.
   
   Hypothesis P_id_U42 : A ->A ->A ->A.
   
   Hypothesis P_id_U22 : A ->A ->A ->A ->A.
   
   Hypothesis P_id_DIN : A ->A.
   
   Hypothesis P_id_U32 : A ->A ->A ->A ->A.
   
   Hypothesis P_id_U41_monotonic :
    forall x8 x10 x9 x7, 
     (A0 <= x10)/\ (x10 <= x9) ->
      (A0 <= x8)/\ (x8 <= x7) ->P_id_U41 x8 x10 <= P_id_U41 x7 x9.
   
   Hypothesis P_id_U21_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (A0 <= x12)/\ (x12 <= x11) ->
      (A0 <= x10)/\ (x10 <= x9) ->
       (A0 <= x8)/\ (x8 <= x7) ->P_id_U21 x8 x10 x12 <= P_id_U21 x7 x9 x11.
   
   Hypothesis P_id_U31_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (A0 <= x12)/\ (x12 <= x11) ->
      (A0 <= x10)/\ (x10 <= x9) ->
       (A0 <= x8)/\ (x8 <= x7) ->P_id_U31 x8 x10 x12 <= P_id_U31 x7 x9 x11.
   
   Hypothesis P_id_U42_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (A0 <= x12)/\ (x12 <= x11) ->
      (A0 <= x10)/\ (x10 <= x9) ->
       (A0 <= x8)/\ (x8 <= x7) ->P_id_U42 x8 x10 x12 <= P_id_U42 x7 x9 x11.
   
   Hypothesis P_id_U22_monotonic :
    forall x8 x12 x10 x14 x9 x13 x11 x7, 
     (A0 <= x14)/\ (x14 <= x13) ->
      (A0 <= x12)/\ (x12 <= x11) ->
       (A0 <= x10)/\ (x10 <= x9) ->
        (A0 <= x8)/\ (x8 <= x7) ->
         P_id_U22 x8 x10 x12 x14 <= P_id_U22 x7 x9 x11 x13.
   
   Hypothesis P_id_DIN_monotonic :
    forall x8 x7, (A0 <= x8)/\ (x8 <= x7) ->P_id_DIN x8 <= P_id_DIN x7.
   
   Hypothesis P_id_U32_monotonic :
    forall x8 x12 x10 x14 x9 x13 x11 x7, 
     (A0 <= x14)/\ (x14 <= x13) ->
      (A0 <= x12)/\ (x12 <= x11) ->
       (A0 <= x10)/\ (x10 <= x9) ->
        (A0 <= x8)/\ (x8 <= x7) ->
         P_id_U32 x8 x10 x12 x14 <= P_id_U32 x7 x9 x11 x13.
   
   Definition marked_measure t := 
     match t with
       | (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil)) =>
        P_id_U41 (measure x8) (measure x7)
       | (algebra.Alg.Term algebra.F.id_u21 (x9::x8::x7::nil)) =>
        P_id_U21 (measure x9) (measure x8) (measure x7)
       | (algebra.Alg.Term algebra.F.id_u31 (x9::x8::x7::nil)) =>
        P_id_U31 (measure x9) (measure x8) (measure x7)
       | (algebra.Alg.Term algebra.F.id_u42 (x9::x8::x7::nil)) =>
        P_id_U42 (measure x9) (measure x8) (measure x7)
       | (algebra.Alg.Term algebra.F.id_u22 (x10::x9::x8::x7::nil)) =>
        P_id_U22 (measure x10) (measure x9) (measure x8) (measure x7)
       | (algebra.Alg.Term algebra.F.id_din (x7::nil)) =>
        P_id_DIN (measure x7)
       | (algebra.Alg.Term algebra.F.id_u32 (x10::x9::x8::x7::nil)) =>
        P_id_U32 (measure x10) (measure x9) (measure x8) (measure x7)
       | _ => 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 x9 x8 x7;apply (P_id_U31 x9 x8 x7).
     
     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 x9 x8 x7;apply (P_id_U21 x9 x8 x7).
     
     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 x10 x9 x8 x7;apply (P_id_U22 x10 x9 x8 x7).
     
     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 x8 x7;apply (P_id_U41 x8 x7).
     
     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 x9 x8 x7;apply (P_id_U42 x9 x8 x7).
     
     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 x10 x9 x8 x7;apply (P_id_U32 x10 x9 x8 x7).
     
     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 x7;apply (P_id_DIN x7).
     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_din =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_u32 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::nil => _
                 | _::_::_::_::nil => _
                 | _::_::_::_::_::_ => _
                 end
              | algebra.F.id_dout =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_plus =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_u42 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::nil => _
                 | _::_::_::_::_ => _
                 end
              | algebra.F.id_times =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_der =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_u41 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_u22 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::nil => _
                 | _::_::_::_::nil => _
                 | _::_::_::_::_::_ => _
                 end
              | algebra.F.id_u21 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::nil => _
                 | _::_::_::_::_ => _
                 end
              | algebra.F.id_u31 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::nil => _
                 | _::_::_::_::_ => _
                 end
              end.
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     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_u41 (x8::x7::nil)) =>
                           P_id_U41 (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u21 (x9::x8::
                             x7::nil)) =>
                           P_id_U21 (measure x9) (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u31 (x9::x8::
                             x7::nil)) =>
                           P_id_U31 (measure x9) (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u42 (x9::x8::
                             x7::nil)) =>
                           P_id_U42 (measure x9) (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u22 (x10::x9::x8::
                             x7::nil)) =>
                           P_id_U22 (measure x10) (measure x9) (measure x8) 
                            (measure x7)
                          | (algebra.Alg.Term algebra.F.id_din (x7::nil)) =>
                           P_id_DIN (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u32 (x10::x9::x8::
                             x7::nil)) =>
                           P_id_U32 (measure x10) (measure x9) (measure x8) 
                            (measure x7)
                          | _ => 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_din_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_u32_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_dout_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_plus_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_u42_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_times_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_der_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_u41_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_u22_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_u21_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_u31_monotonic;assumption.
     clear f.
     intros f.
     case f.
     vm_compute in |-*;intros ;apply P_id_din_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u32_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_dout_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_plus_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u42_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_times_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_der_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u41_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u22_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u21_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u31_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_U31_monotonic;assumption.
     
     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_U21_monotonic;assumption.
     
     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_U22_monotonic;assumption.
     
     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_U41_monotonic;assumption.
     
     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_U42_monotonic;assumption.
     
     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_U32_monotonic;assumption.
     
     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_DIN_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_din : Z ->Z.
   
   Hypothesis P_id_u32 : Z ->Z ->Z ->Z ->Z.
   
   Hypothesis P_id_dout : Z ->Z.
   
   Hypothesis P_id_plus : Z ->Z ->Z.
   
   Hypothesis P_id_u42 : Z ->Z ->Z ->Z.
   
   Hypothesis P_id_times : Z ->Z ->Z.
   
   Hypothesis P_id_der : Z ->Z.
   
   Hypothesis P_id_u41 : Z ->Z ->Z.
   
   Hypothesis P_id_u22 : Z ->Z ->Z ->Z ->Z.
   
   Hypothesis P_id_u21 : Z ->Z ->Z ->Z.
   
   Hypothesis P_id_u31 : Z ->Z ->Z ->Z.
   
   Hypothesis P_id_din_monotonic :
    forall x8 x7, (min_value <= x8)/\ (x8 <= x7) ->P_id_din x8 <= P_id_din x7.
   
   Hypothesis P_id_u32_monotonic :
    forall x8 x12 x10 x14 x9 x13 x11 x7, 
     (min_value <= x14)/\ (x14 <= x13) ->
      (min_value <= x12)/\ (x12 <= x11) ->
       (min_value <= x10)/\ (x10 <= x9) ->
        (min_value <= x8)/\ (x8 <= x7) ->
         P_id_u32 x8 x10 x12 x14 <= P_id_u32 x7 x9 x11 x13.
   
   Hypothesis P_id_dout_monotonic :
    forall x8 x7, 
     (min_value <= x8)/\ (x8 <= x7) ->P_id_dout x8 <= P_id_dout x7.
   
   Hypothesis P_id_plus_monotonic :
    forall x8 x10 x9 x7, 
     (min_value <= x10)/\ (x10 <= x9) ->
      (min_value <= x8)/\ (x8 <= x7) ->P_id_plus x8 x10 <= P_id_plus x7 x9.
   
   Hypothesis P_id_u42_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (min_value <= x12)/\ (x12 <= x11) ->
      (min_value <= x10)/\ (x10 <= x9) ->
       (min_value <= x8)/\ (x8 <= x7) ->
        P_id_u42 x8 x10 x12 <= P_id_u42 x7 x9 x11.
   
   Hypothesis P_id_times_monotonic :
    forall x8 x10 x9 x7, 
     (min_value <= x10)/\ (x10 <= x9) ->
      (min_value <= x8)/\ (x8 <= x7) ->P_id_times x8 x10 <= P_id_times x7 x9.
   
   Hypothesis P_id_der_monotonic :
    forall x8 x7, (min_value <= x8)/\ (x8 <= x7) ->P_id_der x8 <= P_id_der x7.
   
   Hypothesis P_id_u41_monotonic :
    forall x8 x10 x9 x7, 
     (min_value <= x10)/\ (x10 <= x9) ->
      (min_value <= x8)/\ (x8 <= x7) ->P_id_u41 x8 x10 <= P_id_u41 x7 x9.
   
   Hypothesis P_id_u22_monotonic :
    forall x8 x12 x10 x14 x9 x13 x11 x7, 
     (min_value <= x14)/\ (x14 <= x13) ->
      (min_value <= x12)/\ (x12 <= x11) ->
       (min_value <= x10)/\ (x10 <= x9) ->
        (min_value <= x8)/\ (x8 <= x7) ->
         P_id_u22 x8 x10 x12 x14 <= P_id_u22 x7 x9 x11 x13.
   
   Hypothesis P_id_u21_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (min_value <= x12)/\ (x12 <= x11) ->
      (min_value <= x10)/\ (x10 <= x9) ->
       (min_value <= x8)/\ (x8 <= x7) ->
        P_id_u21 x8 x10 x12 <= P_id_u21 x7 x9 x11.
   
   Hypothesis P_id_u31_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (min_value <= x12)/\ (x12 <= x11) ->
      (min_value <= x10)/\ (x10 <= x9) ->
       (min_value <= x8)/\ (x8 <= x7) ->
        P_id_u31 x8 x10 x12 <= P_id_u31 x7 x9 x11.
   
   Hypothesis P_id_din_bounded :
    forall x7, (min_value <= x7) ->min_value <= P_id_din x7.
   
   Hypothesis P_id_u32_bounded :
    forall x8 x10 x9 x7, 
     (min_value <= x7) ->
      (min_value <= x8) ->
       (min_value <= x9) ->
        (min_value <= x10) ->min_value <= P_id_u32 x10 x9 x8 x7.
   
   Hypothesis P_id_dout_bounded :
    forall x7, (min_value <= x7) ->min_value <= P_id_dout x7.
   
   Hypothesis P_id_plus_bounded :
    forall x8 x7, 
     (min_value <= x7) ->(min_value <= x8) ->min_value <= P_id_plus x8 x7.
   
   Hypothesis P_id_u42_bounded :
    forall x8 x9 x7, 
     (min_value <= x7) ->
      (min_value <= x8) ->(min_value <= x9) ->min_value <= P_id_u42 x9 x8 x7.
   
   Hypothesis P_id_times_bounded :
    forall x8 x7, 
     (min_value <= x7) ->(min_value <= x8) ->min_value <= P_id_times x8 x7.
   
   Hypothesis P_id_der_bounded :
    forall x7, (min_value <= x7) ->min_value <= P_id_der x7.
   
   Hypothesis P_id_u41_bounded :
    forall x8 x7, 
     (min_value <= x7) ->(min_value <= x8) ->min_value <= P_id_u41 x8 x7.
   
   Hypothesis P_id_u22_bounded :
    forall x8 x10 x9 x7, 
     (min_value <= x7) ->
      (min_value <= x8) ->
       (min_value <= x9) ->
        (min_value <= x10) ->min_value <= P_id_u22 x10 x9 x8 x7.
   
   Hypothesis P_id_u21_bounded :
    forall x8 x9 x7, 
     (min_value <= x7) ->
      (min_value <= x8) ->(min_value <= x9) ->min_value <= P_id_u21 x9 x8 x7.
   
   Hypothesis P_id_u31_bounded :
    forall x8 x9 x7, 
     (min_value <= x7) ->
      (min_value <= x8) ->(min_value <= x9) ->min_value <= P_id_u31 x9 x8 x7.
   
   Definition measure  := 
     Interp.measure min_value P_id_din P_id_u32 P_id_dout P_id_plus P_id_u42 
      P_id_times P_id_der P_id_u41 P_id_u22 P_id_u21 P_id_u31.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_din (x7::nil)) =>
                    P_id_din (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u32 (x10::x9::x8::
                      x7::nil)) =>
                    P_id_u32 (measure x10) (measure x9) (measure x8) 
                     (measure x7)
                   | (algebra.Alg.Term algebra.F.id_dout (x7::nil)) =>
                    P_id_dout (measure x7)
                   | (algebra.Alg.Term algebra.F.id_plus (x8::x7::nil)) =>
                    P_id_plus (measure x8) (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u42 (x9::x8::x7::nil)) =>
                    P_id_u42 (measure x9) (measure x8) (measure x7)
                   | (algebra.Alg.Term algebra.F.id_times (x8::x7::nil)) =>
                    P_id_times (measure x8) (measure x7)
                   | (algebra.Alg.Term algebra.F.id_der (x7::nil)) =>
                    P_id_der (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil)) =>
                    P_id_u41 (measure x8) (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u22 (x10::x9::x8::
                      x7::nil)) =>
                    P_id_u22 (measure x10) (measure x9) (measure x8) 
                     (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u21 (x9::x8::x7::nil)) =>
                    P_id_u21 (measure x9) (measure x8) (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u31 (x9::x8::x7::nil)) =>
                    P_id_u31 (measure x9) (measure x8) (measure x7)
                   | _ => 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_din_monotonic;assumption.
     intros ;apply P_id_u32_monotonic;assumption.
     intros ;apply P_id_dout_monotonic;assumption.
     intros ;apply P_id_plus_monotonic;assumption.
     intros ;apply P_id_u42_monotonic;assumption.
     intros ;apply P_id_times_monotonic;assumption.
     intros ;apply P_id_der_monotonic;assumption.
     intros ;apply P_id_u41_monotonic;assumption.
     intros ;apply P_id_u22_monotonic;assumption.
     intros ;apply P_id_u21_monotonic;assumption.
     intros ;apply P_id_u31_monotonic;assumption.
     intros ;apply P_id_din_bounded;assumption.
     intros ;apply P_id_u32_bounded;assumption.
     intros ;apply P_id_dout_bounded;assumption.
     intros ;apply P_id_plus_bounded;assumption.
     intros ;apply P_id_u42_bounded;assumption.
     intros ;apply P_id_times_bounded;assumption.
     intros ;apply P_id_der_bounded;assumption.
     intros ;apply P_id_u41_bounded;assumption.
     intros ;apply P_id_u22_bounded;assumption.
     intros ;apply P_id_u21_bounded;assumption.
     intros ;apply P_id_u31_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Hypothesis P_id_U41 : Z ->Z ->Z.
   
   Hypothesis P_id_U21 : Z ->Z ->Z ->Z.
   
   Hypothesis P_id_U31 : Z ->Z ->Z ->Z.
   
   Hypothesis P_id_U42 : Z ->Z ->Z ->Z.
   
   Hypothesis P_id_U22 : Z ->Z ->Z ->Z ->Z.
   
   Hypothesis P_id_DIN : Z ->Z.
   
   Hypothesis P_id_U32 : Z ->Z ->Z ->Z ->Z.
   
   Hypothesis P_id_U41_monotonic :
    forall x8 x10 x9 x7, 
     (min_value <= x10)/\ (x10 <= x9) ->
      (min_value <= x8)/\ (x8 <= x7) ->P_id_U41 x8 x10 <= P_id_U41 x7 x9.
   
   Hypothesis P_id_U21_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (min_value <= x12)/\ (x12 <= x11) ->
      (min_value <= x10)/\ (x10 <= x9) ->
       (min_value <= x8)/\ (x8 <= x7) ->
        P_id_U21 x8 x10 x12 <= P_id_U21 x7 x9 x11.
   
   Hypothesis P_id_U31_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (min_value <= x12)/\ (x12 <= x11) ->
      (min_value <= x10)/\ (x10 <= x9) ->
       (min_value <= x8)/\ (x8 <= x7) ->
        P_id_U31 x8 x10 x12 <= P_id_U31 x7 x9 x11.
   
   Hypothesis P_id_U42_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (min_value <= x12)/\ (x12 <= x11) ->
      (min_value <= x10)/\ (x10 <= x9) ->
       (min_value <= x8)/\ (x8 <= x7) ->
        P_id_U42 x8 x10 x12 <= P_id_U42 x7 x9 x11.
   
   Hypothesis P_id_U22_monotonic :
    forall x8 x12 x10 x14 x9 x13 x11 x7, 
     (min_value <= x14)/\ (x14 <= x13) ->
      (min_value <= x12)/\ (x12 <= x11) ->
       (min_value <= x10)/\ (x10 <= x9) ->
        (min_value <= x8)/\ (x8 <= x7) ->
         P_id_U22 x8 x10 x12 x14 <= P_id_U22 x7 x9 x11 x13.
   
   Hypothesis P_id_DIN_monotonic :
    forall x8 x7, (min_value <= x8)/\ (x8 <= x7) ->P_id_DIN x8 <= P_id_DIN x7.
   
   Hypothesis P_id_U32_monotonic :
    forall x8 x12 x10 x14 x9 x13 x11 x7, 
     (min_value <= x14)/\ (x14 <= x13) ->
      (min_value <= x12)/\ (x12 <= x11) ->
       (min_value <= x10)/\ (x10 <= x9) ->
        (min_value <= x8)/\ (x8 <= x7) ->
         P_id_U32 x8 x10 x12 x14 <= P_id_U32 x7 x9 x11 x13.
   
   Definition marked_measure  := 
     Interp.marked_measure min_value P_id_din P_id_u32 P_id_dout P_id_plus 
      P_id_u42 P_id_times P_id_der P_id_u41 P_id_u22 P_id_u21 P_id_u31 
      P_id_U41 P_id_U21 P_id_U31 P_id_U42 P_id_U22 P_id_DIN P_id_U32.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil)) =>
                           P_id_U41 (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u21 (x9::x8::
                             x7::nil)) =>
                           P_id_U21 (measure x9) (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u31 (x9::x8::
                             x7::nil)) =>
                           P_id_U31 (measure x9) (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u42 (x9::x8::
                             x7::nil)) =>
                           P_id_U42 (measure x9) (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u22 (x10::x9::x8::
                             x7::nil)) =>
                           P_id_U22 (measure x10) (measure x9) (measure x8) 
                            (measure x7)
                          | (algebra.Alg.Term algebra.F.id_din (x7::nil)) =>
                           P_id_DIN (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u32 (x10::x9::x8::
                             x7::nil)) =>
                           P_id_U32 (measure x10) (measure x9) (measure x8) 
                            (measure x7)
                          | _ => 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_din_monotonic;assumption.
     intros ;apply P_id_u32_monotonic;assumption.
     intros ;apply P_id_dout_monotonic;assumption.
     intros ;apply P_id_plus_monotonic;assumption.
     intros ;apply P_id_u42_monotonic;assumption.
     intros ;apply P_id_times_monotonic;assumption.
     intros ;apply P_id_der_monotonic;assumption.
     intros ;apply P_id_u41_monotonic;assumption.
     intros ;apply P_id_u22_monotonic;assumption.
     intros ;apply P_id_u21_monotonic;assumption.
     intros ;apply P_id_u31_monotonic;assumption.
     intros ;apply P_id_din_bounded;assumption.
     intros ;apply P_id_u32_bounded;assumption.
     intros ;apply P_id_dout_bounded;assumption.
     intros ;apply P_id_plus_bounded;assumption.
     intros ;apply P_id_u42_bounded;assumption.
     intros ;apply P_id_times_bounded;assumption.
     intros ;apply P_id_der_bounded;assumption.
     intros ;apply P_id_u41_bounded;assumption.
     intros ;apply P_id_u22_bounded;assumption.
     intros ;apply P_id_u21_bounded;assumption.
     intros ;apply P_id_u31_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_DIN_monotonic;assumption.
     intros ;apply P_id_U32_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 := 
    (* <din(der(plus(X_,Y_))),u21(din(der(X_)),X_,Y_)> *)
   | DP_R_xml_0_0 :
    forall x2 x1 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
        algebra.F.id_plus (x1::x2::nil))::nil)) x7) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_u21 ((algebra.Alg.Term 
                  algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
                  (x1::nil))::nil))::x1::x2::nil)) 
       (algebra.Alg.Term algebra.F.id_din (x7::nil))
    (* <din(der(plus(X_,Y_))),din(der(X_))> *)
   | DP_R_xml_0_1 :
    forall x2 x1 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
        algebra.F.id_plus (x1::x2::nil))::nil)) x7) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_din ((algebra.Alg.Term 
                  algebra.F.id_der (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_din (x7::nil))
    (* <u21(dout(DX_),X_,Y_),u22(din(der(Y_)),X_,Y_,DX_)> *)
   | DP_R_xml_0_2 :
    forall x8 x2 x1 x9 x3 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x9) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x1 x8) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x2 x7) ->
        DP_R_xml_0 (algebra.Alg.Term algebra.F.id_u22 ((algebra.Alg.Term 
                    algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
                    (x2::nil))::nil))::x1::x2::x3::nil)) 
         (algebra.Alg.Term algebra.F.id_u21 (x9::x8::x7::nil))
    (* <u21(dout(DX_),X_,Y_),din(der(Y_))> *)
   | DP_R_xml_0_3 :
    forall x8 x2 x1 x9 x3 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x9) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x1 x8) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x2 x7) ->
        DP_R_xml_0 (algebra.Alg.Term algebra.F.id_din ((algebra.Alg.Term 
                    algebra.F.id_der (x2::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_u21 (x9::x8::x7::nil))
    (* <din(der(times(X_,Y_))),u31(din(der(X_)),X_,Y_)> *)
   | DP_R_xml_0_4 :
    forall x2 x1 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
        algebra.F.id_times (x1::x2::nil))::nil)) x7) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_u31 ((algebra.Alg.Term 
                  algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
                  (x1::nil))::nil))::x1::x2::nil)) 
       (algebra.Alg.Term algebra.F.id_din (x7::nil))
    (* <din(der(times(X_,Y_))),din(der(X_))> *)
   | DP_R_xml_0_5 :
    forall x2 x1 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
        algebra.F.id_times (x1::x2::nil))::nil)) x7) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_din ((algebra.Alg.Term 
                  algebra.F.id_der (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_din (x7::nil))
    (* <u31(dout(DX_),X_,Y_),u32(din(der(Y_)),X_,Y_,DX_)> *)
   | DP_R_xml_0_6 :
    forall x8 x2 x1 x9 x3 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x9) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x1 x8) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x2 x7) ->
        DP_R_xml_0 (algebra.Alg.Term algebra.F.id_u32 ((algebra.Alg.Term 
                    algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
                    (x2::nil))::nil))::x1::x2::x3::nil)) 
         (algebra.Alg.Term algebra.F.id_u31 (x9::x8::x7::nil))
    (* <u31(dout(DX_),X_,Y_),din(der(Y_))> *)
   | DP_R_xml_0_7 :
    forall x8 x2 x1 x9 x3 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x9) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x1 x8) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x2 x7) ->
        DP_R_xml_0 (algebra.Alg.Term algebra.F.id_din ((algebra.Alg.Term 
                    algebra.F.id_der (x2::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_u31 (x9::x8::x7::nil))
    (* <din(der(der(X_))),u41(din(der(X_)),X_)> *)
   | DP_R_xml_0_8 :
    forall x1 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
        algebra.F.id_der (x1::nil))::nil)) x7) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_u41 ((algebra.Alg.Term 
                  algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
                  (x1::nil))::nil))::x1::nil)) 
       (algebra.Alg.Term algebra.F.id_din (x7::nil))
    (* <din(der(der(X_))),din(der(X_))> *)
   | DP_R_xml_0_9 :
    forall x1 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
        algebra.F.id_der (x1::nil))::nil)) x7) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_din ((algebra.Alg.Term 
                  algebra.F.id_der (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_din (x7::nil))
    (* <u41(dout(DX_),X_),u42(din(der(DX_)),X_,DX_)> *)
   | DP_R_xml_0_10 :
    forall x8 x1 x3 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x8) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x1 x7) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_u42 ((algebra.Alg.Term 
                   algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
                   (x3::nil))::nil))::x1::x3::nil)) 
        (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil))
    (* <u41(dout(DX_),X_),din(der(DX_))> *)
   | DP_R_xml_0_11 :
    forall x8 x1 x3 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x8) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x1 x7) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_din ((algebra.Alg.Term 
                   algebra.F.id_der (x3::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_u41 (x8::x7::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_non_scc_1  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <u41(dout(DX_),X_),u42(din(der(DX_)),X_,DX_)> *)
    | DP_R_xml_0_non_scc_1_0 :
     forall x8 x1 x3 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x8) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x1 x7) ->
        DP_R_xml_0_non_scc_1 (algebra.Alg.Term algebra.F.id_u42 
                              ((algebra.Alg.Term algebra.F.id_din 
                              ((algebra.Alg.Term algebra.F.id_der 
                              (x3::nil))::nil))::x1::x3::nil)) 
         (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_1 :
   forall x y, 
    (DP_R_xml_0_non_scc_1 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x.
  Proof.
    intros x y h.
    
    inversion h;clear h;subst;
     constructor;intros _y _h;inversion _h;clear _h;subst;
      (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||
      (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_2  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <u31(dout(DX_),X_,Y_),u32(din(der(Y_)),X_,Y_,DX_)> *)
    | DP_R_xml_0_non_scc_2_0 :
     forall x8 x2 x1 x9 x3 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x9) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x1 x8) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   x2 x7) ->
         DP_R_xml_0_non_scc_2 (algebra.Alg.Term algebra.F.id_u32 
                               ((algebra.Alg.Term algebra.F.id_din 
                               ((algebra.Alg.Term algebra.F.id_der 
                               (x2::nil))::nil))::x1::x2::x3::nil)) 
          (algebra.Alg.Term algebra.F.id_u31 (x9::x8::x7::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_2 :
   forall x y, 
    (DP_R_xml_0_non_scc_2 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x.
  Proof.
    intros x y h.
    
    inversion h;clear h;subst;
     constructor;intros _y _h;inversion _h;clear _h;subst;
      (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||
      (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_3  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <u21(dout(DX_),X_,Y_),u22(din(der(Y_)),X_,Y_,DX_)> *)
    | DP_R_xml_0_non_scc_3_0 :
     forall x8 x2 x1 x9 x3 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x9) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x1 x8) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   x2 x7) ->
         DP_R_xml_0_non_scc_3 (algebra.Alg.Term algebra.F.id_u22 
                               ((algebra.Alg.Term algebra.F.id_din 
                               ((algebra.Alg.Term algebra.F.id_der 
                               (x2::nil))::nil))::x1::x2::x3::nil)) 
          (algebra.Alg.Term algebra.F.id_u21 (x9::x8::x7::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_3 :
   forall x y, 
    (DP_R_xml_0_non_scc_3 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x.
  Proof.
    intros x y h.
    
    inversion h;clear h;subst;
     constructor;intros _y _h;inversion _h;clear _h;subst;
      (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||
      (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )).
  Qed.
  
  
  Inductive DP_R_xml_0_scc_4  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <u21(dout(DX_),X_,Y_),din(der(Y_))> *)
    | DP_R_xml_0_scc_4_0 :
     forall x8 x2 x1 x9 x3 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x9) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x1 x8) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   x2 x7) ->
         DP_R_xml_0_scc_4 (algebra.Alg.Term algebra.F.id_din 
                           ((algebra.Alg.Term algebra.F.id_der 
                           (x2::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_u21 (x9::x8::x7::nil))
    
     (* <din(der(plus(X_,Y_))),u21(din(der(X_)),X_,Y_)> *)
    | DP_R_xml_0_scc_4_1 :
     forall x2 x1 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
         algebra.F.id_plus (x1::x2::nil))::nil)) x7) ->
       DP_R_xml_0_scc_4 (algebra.Alg.Term algebra.F.id_u21 
                         ((algebra.Alg.Term algebra.F.id_din 
                         ((algebra.Alg.Term algebra.F.id_der 
                         (x1::nil))::nil))::x1::x2::nil)) 
        (algebra.Alg.Term algebra.F.id_din (x7::nil))
     (* <din(der(plus(X_,Y_))),din(der(X_))> *)
    | DP_R_xml_0_scc_4_2 :
     forall x2 x1 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
         algebra.F.id_plus (x1::x2::nil))::nil)) x7) ->
       DP_R_xml_0_scc_4 (algebra.Alg.Term algebra.F.id_din 
                         ((algebra.Alg.Term algebra.F.id_der 
                         (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_din (x7::nil))
    
     (* <din(der(times(X_,Y_))),u31(din(der(X_)),X_,Y_)> *)
    | DP_R_xml_0_scc_4_3 :
     forall x2 x1 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
         algebra.F.id_times (x1::x2::nil))::nil)) x7) ->
       DP_R_xml_0_scc_4 (algebra.Alg.Term algebra.F.id_u31 
                         ((algebra.Alg.Term algebra.F.id_din 
                         ((algebra.Alg.Term algebra.F.id_der 
                         (x1::nil))::nil))::x1::x2::nil)) 
        (algebra.Alg.Term algebra.F.id_din (x7::nil))
     (* <u31(dout(DX_),X_,Y_),din(der(Y_))> *)
    | DP_R_xml_0_scc_4_4 :
     forall x8 x2 x1 x9 x3 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x9) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x1 x8) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   x2 x7) ->
         DP_R_xml_0_scc_4 (algebra.Alg.Term algebra.F.id_din 
                           ((algebra.Alg.Term algebra.F.id_der 
                           (x2::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_u31 (x9::x8::x7::nil))
     (* <din(der(times(X_,Y_))),din(der(X_))> *)
    | DP_R_xml_0_scc_4_5 :
     forall x2 x1 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
         algebra.F.id_times (x1::x2::nil))::nil)) x7) ->
       DP_R_xml_0_scc_4 (algebra.Alg.Term algebra.F.id_din 
                         ((algebra.Alg.Term algebra.F.id_der 
                         (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_din (x7::nil))
     (* <din(der(der(X_))),u41(din(der(X_)),X_)> *)
    | DP_R_xml_0_scc_4_6 :
     forall x1 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
         algebra.F.id_der (x1::nil))::nil)) x7) ->
       DP_R_xml_0_scc_4 (algebra.Alg.Term algebra.F.id_u41 
                         ((algebra.Alg.Term algebra.F.id_din 
                         ((algebra.Alg.Term algebra.F.id_der 
                         (x1::nil))::nil))::x1::nil)) 
        (algebra.Alg.Term algebra.F.id_din (x7::nil))
     (* <u41(dout(DX_),X_),din(der(DX_))> *)
    | DP_R_xml_0_scc_4_7 :
     forall x8 x1 x3 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x8) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x1 x7) ->
        DP_R_xml_0_scc_4 (algebra.Alg.Term algebra.F.id_din 
                          ((algebra.Alg.Term algebra.F.id_der 
                          (x3::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil))
     (* <din(der(der(X_))),din(der(X_))> *)
    | DP_R_xml_0_scc_4_8 :
     forall x1 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
         algebra.F.id_der (x1::nil))::nil)) x7) ->
       DP_R_xml_0_scc_4 (algebra.Alg.Term algebra.F.id_din 
                         ((algebra.Alg.Term algebra.F.id_der 
                         (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_din (x7::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_4.
   Inductive DP_R_xml_0_scc_4_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <din(der(plus(X_,Y_))),u21(din(der(X_)),X_,Y_)> *)
     | DP_R_xml_0_scc_4_large_0 :
      forall x2 x1 x7, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
          algebra.F.id_plus (x1::x2::nil))::nil)) x7) ->
        DP_R_xml_0_scc_4_large (algebra.Alg.Term algebra.F.id_u21 
                                ((algebra.Alg.Term algebra.F.id_din 
                                ((algebra.Alg.Term algebra.F.id_der 
                                (x1::nil))::nil))::x1::x2::nil)) 
         (algebra.Alg.Term algebra.F.id_din (x7::nil))
      (* <din(der(plus(X_,Y_))),din(der(X_))> *)
     | DP_R_xml_0_scc_4_large_1 :
      forall x2 x1 x7, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
          algebra.F.id_plus (x1::x2::nil))::nil)) x7) ->
        DP_R_xml_0_scc_4_large (algebra.Alg.Term algebra.F.id_din 
                                ((algebra.Alg.Term algebra.F.id_der 
                                (x1::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_din (x7::nil))
     
      (* <din(der(times(X_,Y_))),u31(din(der(X_)),X_,Y_)> *)
     | DP_R_xml_0_scc_4_large_2 :
      forall x2 x1 x7, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
          algebra.F.id_times (x1::x2::nil))::nil)) x7) ->
        DP_R_xml_0_scc_4_large (algebra.Alg.Term algebra.F.id_u31 
                                ((algebra.Alg.Term algebra.F.id_din 
                                ((algebra.Alg.Term algebra.F.id_der 
                                (x1::nil))::nil))::x1::x2::nil)) 
         (algebra.Alg.Term algebra.F.id_din (x7::nil))
      (* <u31(dout(DX_),X_,Y_),din(der(Y_))> *)
     | DP_R_xml_0_scc_4_large_3 :
      forall x8 x2 x1 x9 x3 x7, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x9) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   x1 x8) ->
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    x2 x7) ->
          DP_R_xml_0_scc_4_large (algebra.Alg.Term algebra.F.id_din 
                                  ((algebra.Alg.Term algebra.F.id_der 
                                  (x2::nil))::nil)) 
           (algebra.Alg.Term algebra.F.id_u31 (x9::x8::x7::nil))
      (* <din(der(times(X_,Y_))),din(der(X_))> *)
     | DP_R_xml_0_scc_4_large_4 :
      forall x2 x1 x7, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
          algebra.F.id_times (x1::x2::nil))::nil)) x7) ->
        DP_R_xml_0_scc_4_large (algebra.Alg.Term algebra.F.id_din 
                                ((algebra.Alg.Term algebra.F.id_der 
                                (x1::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_din (x7::nil))
     
      (* <din(der(der(X_))),u41(din(der(X_)),X_)> *)
     | DP_R_xml_0_scc_4_large_5 :
      forall x1 x7, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
          algebra.F.id_der (x1::nil))::nil)) x7) ->
        DP_R_xml_0_scc_4_large (algebra.Alg.Term algebra.F.id_u41 
                                ((algebra.Alg.Term algebra.F.id_din 
                                ((algebra.Alg.Term algebra.F.id_der 
                                (x1::nil))::nil))::x1::nil)) 
         (algebra.Alg.Term algebra.F.id_din (x7::nil))
      (* <u41(dout(DX_),X_),din(der(DX_))> *)
     | DP_R_xml_0_scc_4_large_6 :
      forall x8 x1 x3 x7, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x8) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   x1 x7) ->
         DP_R_xml_0_scc_4_large (algebra.Alg.Term algebra.F.id_din 
                                 ((algebra.Alg.Term algebra.F.id_der 
                                 (x3::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil))
      (* <din(der(der(X_))),din(der(X_))> *)
     | DP_R_xml_0_scc_4_large_7 :
      forall x1 x7, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
          algebra.F.id_der (x1::nil))::nil)) x7) ->
        DP_R_xml_0_scc_4_large (algebra.Alg.Term algebra.F.id_din 
                                ((algebra.Alg.Term algebra.F.id_der 
                                (x1::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_din (x7::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_4_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <u21(dout(DX_),X_,Y_),din(der(Y_))> *)
     | DP_R_xml_0_scc_4_strict_0 :
      forall x8 x2 x1 x9 x3 x7, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x9) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   x1 x8) ->
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    x2 x7) ->
          DP_R_xml_0_scc_4_strict (algebra.Alg.Term algebra.F.id_din 
                                   ((algebra.Alg.Term algebra.F.id_der 
                                   (x2::nil))::nil)) 
           (algebra.Alg.Term algebra.F.id_u21 (x9::x8::x7::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_4_large.
    Inductive DP_R_xml_0_scc_4_large_non_scc_1  :
     algebra.Alg.term ->algebra.Alg.term ->Prop := 
       (* <din(der(plus(X_,Y_))),u21(din(der(X_)),X_,Y_)> *)
      | DP_R_xml_0_scc_4_large_non_scc_1_0 :
       forall x2 x1 x7, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
           algebra.F.id_plus (x1::x2::nil))::nil)) x7) ->
         DP_R_xml_0_scc_4_large_non_scc_1 (algebra.Alg.Term algebra.F.id_u21 
                                           ((algebra.Alg.Term 
                                           algebra.F.id_din 
                                           ((algebra.Alg.Term 
                                           algebra.F.id_der 
                                           (x1::nil))::nil))::x1::x2::nil)) 
          (algebra.Alg.Term algebra.F.id_din (x7::nil))
    .
    
    
    Lemma acc_DP_R_xml_0_scc_4_large_non_scc_1 :
     forall x y, 
      (DP_R_xml_0_scc_4_large_non_scc_1 x y) ->
       Acc WF_DP_R_xml_0_scc_4.DP_R_xml_0_scc_4_large x.
    Proof.
      intros x y h.
      
      inversion h;clear h;subst;
       constructor;intros _y _h;inversion _h;clear _h;subst;
        (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||
        (eapply Hrec;
          econstructor (eassumption)||(algebra.Alg_ext.star_refl' )).
    Qed.
    
    
    Inductive DP_R_xml_0_scc_4_large_scc_2  :
     algebra.Alg.term ->algebra.Alg.term ->Prop := 
       (* <din(der(times(X_,Y_))),u31(din(der(X_)),X_,Y_)> *)
      | DP_R_xml_0_scc_4_large_scc_2_0 :
       forall x2 x1 x7, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
           algebra.F.id_times (x1::x2::nil))::nil)) x7) ->
         DP_R_xml_0_scc_4_large_scc_2 (algebra.Alg.Term algebra.F.id_u31 
                                       ((algebra.Alg.Term algebra.F.id_din 
                                       ((algebra.Alg.Term algebra.F.id_der 
                                       (x1::nil))::nil))::x1::x2::nil)) 
          (algebra.Alg.Term algebra.F.id_din (x7::nil))
      
       (* <u31(dout(DX_),X_,Y_),din(der(Y_))> *)
      | DP_R_xml_0_scc_4_large_scc_2_1 :
       forall x8 x2 x1 x9 x3 x7, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x9) ->
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    x1 x8) ->
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     x2 x7) ->
           DP_R_xml_0_scc_4_large_scc_2 (algebra.Alg.Term algebra.F.id_din 
                                         ((algebra.Alg.Term algebra.F.id_der 
                                         (x2::nil))::nil)) 
            (algebra.Alg.Term algebra.F.id_u31 (x9::x8::x7::nil))
      
       (* <din(der(plus(X_,Y_))),din(der(X_))> *)
      | DP_R_xml_0_scc_4_large_scc_2_2 :
       forall x2 x1 x7, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
           algebra.F.id_plus (x1::x2::nil))::nil)) x7) ->
         DP_R_xml_0_scc_4_large_scc_2 (algebra.Alg.Term algebra.F.id_din 
                                       ((algebra.Alg.Term algebra.F.id_der 
                                       (x1::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_din (x7::nil))
      
       (* <din(der(times(X_,Y_))),din(der(X_))> *)
      | DP_R_xml_0_scc_4_large_scc_2_3 :
       forall x2 x1 x7, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
           algebra.F.id_times (x1::x2::nil))::nil)) x7) ->
         DP_R_xml_0_scc_4_large_scc_2 (algebra.Alg.Term algebra.F.id_din 
                                       ((algebra.Alg.Term algebra.F.id_der 
                                       (x1::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_din (x7::nil))
      
       (* <din(der(der(X_))),u41(din(der(X_)),X_)> *)
      | DP_R_xml_0_scc_4_large_scc_2_4 :
       forall x1 x7, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
           algebra.F.id_der (x1::nil))::nil)) x7) ->
         DP_R_xml_0_scc_4_large_scc_2 (algebra.Alg.Term algebra.F.id_u41 
                                       ((algebra.Alg.Term algebra.F.id_din 
                                       ((algebra.Alg.Term algebra.F.id_der 
                                       (x1::nil))::nil))::x1::nil)) 
          (algebra.Alg.Term algebra.F.id_din (x7::nil))
      
       (* <u41(dout(DX_),X_),din(der(DX_))> *)
      | DP_R_xml_0_scc_4_large_scc_2_5 :
       forall x8 x1 x3 x7, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x8) ->
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    x1 x7) ->
          DP_R_xml_0_scc_4_large_scc_2 (algebra.Alg.Term algebra.F.id_din 
                                        ((algebra.Alg.Term algebra.F.id_der 
                                        (x3::nil))::nil)) 
           (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil))
      
       (* <din(der(der(X_))),din(der(X_))> *)
      | DP_R_xml_0_scc_4_large_scc_2_6 :
       forall x1 x7, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
           algebra.F.id_der (x1::nil))::nil)) x7) ->
         DP_R_xml_0_scc_4_large_scc_2 (algebra.Alg.Term algebra.F.id_din 
                                       ((algebra.Alg.Term algebra.F.id_der 
                                       (x1::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_din (x7::nil))
    .
    
    
    Module WF_DP_R_xml_0_scc_4_large_scc_2.
     Inductive DP_R_xml_0_scc_4_large_scc_2_large  :
      algebra.Alg.term ->algebra.Alg.term ->Prop := 
        (* <din(der(times(X_,Y_))),u31(din(der(X_)),X_,Y_)> *)
       | DP_R_xml_0_scc_4_large_scc_2_large_0 :
        forall x2 x1 x7, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
            algebra.F.id_times (x1::x2::nil))::nil)) x7) ->
          DP_R_xml_0_scc_4_large_scc_2_large (algebra.Alg.Term 
                                              algebra.F.id_u31 
                                              ((algebra.Alg.Term 
                                              algebra.F.id_din 
                                              ((algebra.Alg.Term 
                                              algebra.F.id_der 
                                              (x1::nil))::nil))::x1::
                                              x2::nil)) 
           (algebra.Alg.Term algebra.F.id_din (x7::nil))
       
        (* <din(der(plus(X_,Y_))),din(der(X_))> *)
       | DP_R_xml_0_scc_4_large_scc_2_large_1 :
        forall x2 x1 x7, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
            algebra.F.id_plus (x1::x2::nil))::nil)) x7) ->
          DP_R_xml_0_scc_4_large_scc_2_large (algebra.Alg.Term 
                                              algebra.F.id_din 
                                              ((algebra.Alg.Term 
                                              algebra.F.id_der 
                                              (x1::nil))::nil)) 
           (algebra.Alg.Term algebra.F.id_din (x7::nil))
       
        (* <din(der(times(X_,Y_))),din(der(X_))> *)
       | DP_R_xml_0_scc_4_large_scc_2_large_2 :
        forall x2 x1 x7, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
            algebra.F.id_times (x1::x2::nil))::nil)) x7) ->
          DP_R_xml_0_scc_4_large_scc_2_large (algebra.Alg.Term 
                                              algebra.F.id_din 
                                              ((algebra.Alg.Term 
                                              algebra.F.id_der 
                                              (x1::nil))::nil)) 
           (algebra.Alg.Term algebra.F.id_din (x7::nil))
       
        (* <din(der(der(X_))),u41(din(der(X_)),X_)> *)
       | DP_R_xml_0_scc_4_large_scc_2_large_3 :
        forall x1 x7, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
            algebra.F.id_der (x1::nil))::nil)) x7) ->
          DP_R_xml_0_scc_4_large_scc_2_large (algebra.Alg.Term 
                                              algebra.F.id_u41 
                                              ((algebra.Alg.Term 
                                              algebra.F.id_din 
                                              ((algebra.Alg.Term 
                                              algebra.F.id_der 
                                              (x1::nil))::nil))::x1::nil)) 
           (algebra.Alg.Term algebra.F.id_din (x7::nil))
       
        (* <u41(dout(DX_),X_),din(der(DX_))> *)
       | DP_R_xml_0_scc_4_large_scc_2_large_4 :
        forall x8 x1 x3 x7, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x8) ->
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     x1 x7) ->
           DP_R_xml_0_scc_4_large_scc_2_large (algebra.Alg.Term 
                                               algebra.F.id_din 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_der 
                                               (x3::nil))::nil)) 
            (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil))
       
        (* <din(der(der(X_))),din(der(X_))> *)
       | DP_R_xml_0_scc_4_large_scc_2_large_5 :
        forall x1 x7, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
            algebra.F.id_der (x1::nil))::nil)) x7) ->
          DP_R_xml_0_scc_4_large_scc_2_large (algebra.Alg.Term 
                                              algebra.F.id_din 
                                              ((algebra.Alg.Term 
                                              algebra.F.id_der 
                                              (x1::nil))::nil)) 
           (algebra.Alg.Term algebra.F.id_din (x7::nil))
     .
     
     
     Inductive DP_R_xml_0_scc_4_large_scc_2_strict  :
      algebra.Alg.term ->algebra.Alg.term ->Prop := 
        (* <u31(dout(DX_),X_,Y_),din(der(Y_))> *)
       | DP_R_xml_0_scc_4_large_scc_2_strict_0 :
        forall x8 x2 x1 x9 x3 x7, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x9) ->
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     x1 x8) ->
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                      x2 x7) ->
            DP_R_xml_0_scc_4_large_scc_2_strict (algebra.Alg.Term 
                                                 algebra.F.id_din 
                                                 ((algebra.Alg.Term 
                                                 algebra.F.id_der 
                                                 (x2::nil))::nil)) 
             (algebra.Alg.Term algebra.F.id_u31 (x9::x8::x7::nil))
     .
     
     
     Module WF_DP_R_xml_0_scc_4_large_scc_2_large.
      Inductive DP_R_xml_0_scc_4_large_scc_2_large_non_scc_1  :
       algebra.Alg.term ->algebra.Alg.term ->Prop := 
         (* <din(der(times(X_,Y_))),u31(din(der(X_)),X_,Y_)> *)
        | DP_R_xml_0_scc_4_large_scc_2_large_non_scc_1_0 :
         forall x2 x1 x7, 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     
            (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
             algebra.F.id_times (x1::x2::nil))::nil)) x7) ->
           DP_R_xml_0_scc_4_large_scc_2_large_non_scc_1 (algebra.Alg.Term 
                                                         algebra.F.id_u31 
                                                         ((algebra.Alg.Term 
                                                         algebra.F.id_din 
                                                         ((algebra.Alg.Term 
                                                         algebra.F.id_der 
                                                         (x1::nil))::nil))::
                                                         x1::x2::nil)) 
            (algebra.Alg.Term algebra.F.id_din (x7::nil))
      .
      
      
      Lemma acc_DP_R_xml_0_scc_4_large_scc_2_large_non_scc_1 :
       forall x y, 
        (DP_R_xml_0_scc_4_large_scc_2_large_non_scc_1 x y) ->
         Acc WF_DP_R_xml_0_scc_4_large_scc_2.DP_R_xml_0_scc_4_large_scc_2_large 
          x.
      Proof.
        intros x y h.
        
        inversion h;clear h;subst;
         constructor;intros _y _h;inversion _h;clear _h;subst;
          (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||
          (eapply Hrec;
            econstructor (eassumption)||(algebra.Alg_ext.star_refl' )).
      Qed.
      
      
      Inductive DP_R_xml_0_scc_4_large_scc_2_large_scc_2  :
       algebra.Alg.term ->algebra.Alg.term ->Prop := 
         (* <din(der(times(X_,Y_))),din(der(X_))> *)
        | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_0 :
         forall x2 x1 x7, 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     
            (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
             algebra.F.id_times (x1::x2::nil))::nil)) x7) ->
           DP_R_xml_0_scc_4_large_scc_2_large_scc_2 (algebra.Alg.Term 
                                                     algebra.F.id_din 
                                                     ((algebra.Alg.Term 
                                                     algebra.F.id_der 
                                                     (x1::nil))::nil)) 
            (algebra.Alg.Term algebra.F.id_din (x7::nil))
        
         (* <din(der(plus(X_,Y_))),din(der(X_))> *)
        | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_1 :
         forall x2 x1 x7, 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     
            (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
             algebra.F.id_plus (x1::x2::nil))::nil)) x7) ->
           DP_R_xml_0_scc_4_large_scc_2_large_scc_2 (algebra.Alg.Term 
                                                     algebra.F.id_din 
                                                     ((algebra.Alg.Term 
                                                     algebra.F.id_der 
                                                     (x1::nil))::nil)) 
            (algebra.Alg.Term algebra.F.id_din (x7::nil))
        
         (* <din(der(der(X_))),u41(din(der(X_)),X_)> *)
        | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_2 :
         forall x1 x7, 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     
            (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
             algebra.F.id_der (x1::nil))::nil)) x7) ->
           DP_R_xml_0_scc_4_large_scc_2_large_scc_2 (algebra.Alg.Term 
                                                     algebra.F.id_u41 
                                                     ((algebra.Alg.Term 
                                                     algebra.F.id_din 
                                                     ((algebra.Alg.Term 
                                                     algebra.F.id_der 
                                                     (x1::nil))::nil))::
                                                     x1::nil)) 
            (algebra.Alg.Term algebra.F.id_din (x7::nil))
        
         (* <u41(dout(DX_),X_),din(der(DX_))> *)
        | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_3 :
         forall x8 x1 x3 x7, 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     
            (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x8) ->
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                      x1 x7) ->
            DP_R_xml_0_scc_4_large_scc_2_large_scc_2 (algebra.Alg.Term 
                                                      algebra.F.id_din 
                                                      ((algebra.Alg.Term 
                                                      algebra.F.id_der 
                                                      (x3::nil))::nil)) 
             (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil))
        
         (* <din(der(der(X_))),din(der(X_))> *)
        | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_4 :
         forall x1 x7, 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     
            (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
             algebra.F.id_der (x1::nil))::nil)) x7) ->
           DP_R_xml_0_scc_4_large_scc_2_large_scc_2 (algebra.Alg.Term 
                                                     algebra.F.id_din 
                                                     ((algebra.Alg.Term 
                                                     algebra.F.id_der 
                                                     (x1::nil))::nil)) 
            (algebra.Alg.Term algebra.F.id_din (x7::nil))
      .
      
      
      Module WF_DP_R_xml_0_scc_4_large_scc_2_large_scc_2.
       Inductive DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large  :
        algebra.Alg.term ->algebra.Alg.term ->Prop := 
          (* <din(der(times(X_,Y_))),din(der(X_))> *)
         | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_0 :
          forall x2 x1 x7, 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                      
             (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
              algebra.F.id_times (x1::x2::nil))::nil)) x7) ->
            DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large (algebra.Alg.Term 
                                                            algebra.F.id_din 
                                                            ((algebra.Alg.Term 
                                                            algebra.F.id_der 
                                                            (x1::nil))::nil)) 
             (algebra.Alg.Term algebra.F.id_din (x7::nil))
         
          (* <din(der(plus(X_,Y_))),din(der(X_))> *)
         | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_1 :
          forall x2 x1 x7, 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                      
             (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
              algebra.F.id_plus (x1::x2::nil))::nil)) x7) ->
            DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large (algebra.Alg.Term 
                                                            algebra.F.id_din 
                                                            ((algebra.Alg.Term 
                                                            algebra.F.id_der 
                                                            (x1::nil))::nil)) 
             (algebra.Alg.Term algebra.F.id_din (x7::nil))
         
          (* <din(der(der(X_))),u41(din(der(X_)),X_)> *)
         | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_2 :
          forall x1 x7, 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                      
             (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
              algebra.F.id_der (x1::nil))::nil)) x7) ->
            DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large (algebra.Alg.Term 
                                                            algebra.F.id_u41 
                                                            ((algebra.Alg.Term 
                                                            algebra.F.id_din 
                                                            ((algebra.Alg.Term 
                                                            algebra.F.id_der 
                                                            (x1::nil))::nil))::
                                                            x1::nil)) 
             (algebra.Alg.Term algebra.F.id_din (x7::nil))
         
          (* <din(der(der(X_))),din(der(X_))> *)
         | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_3 :
          forall x1 x7, 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                      
             (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
              algebra.F.id_der (x1::nil))::nil)) x7) ->
            DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large (algebra.Alg.Term 
                                                            algebra.F.id_din 
                                                            ((algebra.Alg.Term 
                                                            algebra.F.id_der 
                                                            (x1::nil))::nil)) 
             (algebra.Alg.Term algebra.F.id_din (x7::nil))
       .
       
       
       Inductive DP_R_xml_0_scc_4_large_scc_2_large_scc_2_strict  :
        algebra.Alg.term ->algebra.Alg.term ->Prop := 
          (* <u41(dout(DX_),X_),din(der(DX_))> *)
         | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_strict_0 :
          forall x8 x1 x3 x7, 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                      
             (algebra.Alg.Term algebra.F.id_dout (x3::nil)) x8) ->
            (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                       x1 x7) ->
             DP_R_xml_0_scc_4_large_scc_2_large_scc_2_strict (algebra.Alg.Term 
                                                              algebra.F.id_din 
                                                              ((algebra.Alg.Term 
                                                              algebra.F.id_der 
                                                              (x3::nil))::nil))
                                                               
              (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil))
       .
       
       
       Module WF_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large.
        Inductive DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_non_scc_1  :
         algebra.Alg.term ->algebra.Alg.term ->Prop := 
           (* <din(der(der(X_))),u41(din(der(X_)),X_)> *)
          | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_non_scc_1_0 :
           forall x1 x7, 
            (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                       
              (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
               algebra.F.id_der (x1::nil))::nil)) x7) ->
             DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_non_scc_1 
              (algebra.Alg.Term algebra.F.id_u41 ((algebra.Alg.Term 
               algebra.F.id_din ((algebra.Alg.Term algebra.F.id_der 
               (x1::nil))::nil))::x1::nil)) 
              (algebra.Alg.Term algebra.F.id_din (x7::nil))
        .
        
        
        Lemma acc_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_non_scc_1 :
         forall x y, 
          (DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_non_scc_1 x y) ->
           Acc WF_DP_R_xml_0_scc_4_large_scc_2_large_scc_2.DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large 
            x.
        Proof.
          intros x y h.
          
          inversion h;clear h;subst;
           constructor;intros _y _h;inversion _h;clear _h;subst;
            (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||
            (eapply Hrec;
              econstructor (eassumption)||(algebra.Alg_ext.star_refl' )).
        Qed.
        
        
        Inductive DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2  :
         algebra.Alg.term ->algebra.Alg.term ->Prop := 
           (* <din(der(plus(X_,Y_))),din(der(X_))> *)
          | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_0 :
           forall x2 x1 x7, 
            (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                       
              (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
               algebra.F.id_plus (x1::x2::nil))::nil)) x7) ->
             DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2 (algebra.Alg.Term 
                                                                   algebra.F.id_din 
                                                                   ((
                                                                   algebra.Alg.Term 
                                                                   algebra.F.id_der 
                                                                   (x1::nil))::nil))
                                                                    
              (algebra.Alg.Term algebra.F.id_din (x7::nil))
          
           (* <din(der(times(X_,Y_))),din(der(X_))> *)
          | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_1 :
           forall x2 x1 x7, 
            (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                       
              (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
               algebra.F.id_times (x1::x2::nil))::nil)) x7) ->
             DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2 (algebra.Alg.Term 
                                                                   algebra.F.id_din 
                                                                   ((
                                                                   algebra.Alg.Term 
                                                                   algebra.F.id_der 
                                                                   (x1::nil))::nil))
                                                                    
              (algebra.Alg.Term algebra.F.id_din (x7::nil))
          
           (* <din(der(der(X_))),din(der(X_))> *)
          | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_2 :
           forall x1 x7, 
            (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                       
              (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
               algebra.F.id_der (x1::nil))::nil)) x7) ->
             DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2 (algebra.Alg.Term 
                                                                   algebra.F.id_din 
                                                                   ((
                                                                   algebra.Alg.Term 
                                                                   algebra.F.id_der 
                                                                   (x1::nil))::nil))
                                                                    
              (algebra.Alg.Term algebra.F.id_din (x7::nil))
        .
        
        
        Module WF_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2.
         Inductive DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_large  :
          algebra.Alg.term ->algebra.Alg.term ->Prop := 
            (* <din(der(der(X_))),din(der(X_))> *)
           | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_large_0 :
            forall x1 x7, 
             (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                        
               (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
                algebra.F.id_der (x1::nil))::nil)) x7) ->
              DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_large 
               (algebra.Alg.Term algebra.F.id_din ((algebra.Alg.Term 
                algebra.F.id_der (x1::nil))::nil)) 
               (algebra.Alg.Term algebra.F.id_din (x7::nil))
         .
         
         
         Inductive DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_strict  :
          algebra.Alg.term ->algebra.Alg.term ->Prop := 
            (* <din(der(plus(X_,Y_))),din(der(X_))> *)
           | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_strict_0 :
            forall x2 x1 x7, 
             (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                        
               (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
                algebra.F.id_plus (x1::x2::nil))::nil)) x7) ->
              DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_strict 
               (algebra.Alg.Term algebra.F.id_din ((algebra.Alg.Term 
                algebra.F.id_der (x1::nil))::nil)) 
               (algebra.Alg.Term algebra.F.id_din (x7::nil))
           
            (* <din(der(times(X_,Y_))),din(der(X_))> *)
           | DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_strict_1 :
            forall x2 x1 x7, 
             (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                        
               (algebra.Alg.Term algebra.F.id_der ((algebra.Alg.Term 
                algebra.F.id_times (x1::x2::nil))::nil)) x7) ->
              DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_strict 
               (algebra.Alg.Term algebra.F.id_din ((algebra.Alg.Term 
                algebra.F.id_der (x1::nil))::nil)) 
               (algebra.Alg.Term algebra.F.id_din (x7::nil))
         .
         
         
         Module WF_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_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_din (x7:Z) := 2.
          
          Definition P_id_u32 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 0.
          
          Definition P_id_dout (x7:Z) := 1* x7.
          
          Definition P_id_plus (x7:Z) (x8:Z) := 0.
          
          Definition P_id_u42 (x7:Z) (x8:Z) (x9:Z) := 1* x7.
          
          Definition P_id_times (x7:Z) (x8:Z) := 0.
          
          Definition P_id_der (x7:Z) := 1 + 1* x7.
          
          Definition P_id_u41 (x7:Z) (x8:Z) := 2.
          
          Definition P_id_u22 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 0.
          
          Definition P_id_u21 (x7:Z) (x8:Z) (x9:Z) := 2.
          
          Definition P_id_u31 (x7:Z) (x8:Z) (x9:Z) := 1* x7.
          
          Lemma P_id_din_monotonic :
           forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_din x8 <= P_id_din x7.
          Proof.
            intros x8 x7.
            intros [H_1 H_0].
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_u32_monotonic :
           forall x8 x12 x10 x14 x9 x13 x11 x7, 
            (0 <= x14)/\ (x14 <= x13) ->
             (0 <= x12)/\ (x12 <= x11) ->
              (0 <= x10)/\ (x10 <= x9) ->
               (0 <= x8)/\ (x8 <= x7) ->
                P_id_u32 x8 x10 x12 x14 <= P_id_u32 x7 x9 x11 x13.
          Proof.
            intros x14 x13 x12 x11 x10 x9 x8 x7.
            intros [H_1 H_0].
            intros [H_3 H_2].
            intros [H_5 H_4].
            intros [H_7 H_6].
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_dout_monotonic :
           forall x8 x7, 
            (0 <= x8)/\ (x8 <= x7) ->P_id_dout x8 <= P_id_dout x7.
          Proof.
            intros x8 x7.
            intros [H_1 H_0].
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_plus_monotonic :
           forall x8 x10 x9 x7, 
            (0 <= x10)/\ (x10 <= x9) ->
             (0 <= x8)/\ (x8 <= x7) ->P_id_plus x8 x10 <= P_id_plus x7 x9.
          Proof.
            intros x10 x9 x8 x7.
            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_u42_monotonic :
           forall x8 x12 x10 x9 x11 x7, 
            (0 <= x12)/\ (x12 <= x11) ->
             (0 <= x10)/\ (x10 <= x9) ->
              (0 <= x8)/\ (x8 <= x7) ->
               P_id_u42 x8 x10 x12 <= P_id_u42 x7 x9 x11.
          Proof.
            intros x12 x11 x10 x9 x8 x7.
            intros [H_1 H_0].
            intros [H_3 H_2].
            intros [H_5 H_4].
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_times_monotonic :
           forall x8 x10 x9 x7, 
            (0 <= x10)/\ (x10 <= x9) ->
             (0 <= x8)/\ (x8 <= x7) ->P_id_times x8 x10 <= P_id_times x7 x9.
          Proof.
            intros x10 x9 x8 x7.
            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_der_monotonic :
           forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_der x8 <= P_id_der x7.
          Proof.
            intros x8 x7.
            intros [H_1 H_0].
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_u41_monotonic :
           forall x8 x10 x9 x7, 
            (0 <= x10)/\ (x10 <= x9) ->
             (0 <= x8)/\ (x8 <= x7) ->P_id_u41 x8 x10 <= P_id_u41 x7 x9.
          Proof.
            intros x10 x9 x8 x7.
            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_u22_monotonic :
           forall x8 x12 x10 x14 x9 x13 x11 x7, 
            (0 <= x14)/\ (x14 <= x13) ->
             (0 <= x12)/\ (x12 <= x11) ->
              (0 <= x10)/\ (x10 <= x9) ->
               (0 <= x8)/\ (x8 <= x7) ->
                P_id_u22 x8 x10 x12 x14 <= P_id_u22 x7 x9 x11 x13.
          Proof.
            intros x14 x13 x12 x11 x10 x9 x8 x7.
            intros [H_1 H_0].
            intros [H_3 H_2].
            intros [H_5 H_4].
            intros [H_7 H_6].
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_u21_monotonic :
           forall x8 x12 x10 x9 x11 x7, 
            (0 <= x12)/\ (x12 <= x11) ->
             (0 <= x10)/\ (x10 <= x9) ->
              (0 <= x8)/\ (x8 <= x7) ->
               P_id_u21 x8 x10 x12 <= P_id_u21 x7 x9 x11.
          Proof.
            intros x12 x11 x10 x9 x8 x7.
            intros [H_1 H_0].
            intros [H_3 H_2].
            intros [H_5 H_4].
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_u31_monotonic :
           forall x8 x12 x10 x9 x11 x7, 
            (0 <= x12)/\ (x12 <= x11) ->
             (0 <= x10)/\ (x10 <= x9) ->
              (0 <= x8)/\ (x8 <= x7) ->
               P_id_u31 x8 x10 x12 <= P_id_u31 x7 x9 x11.
          Proof.
            intros x12 x11 x10 x9 x8 x7.
            intros [H_1 H_0].
            intros [H_3 H_2].
            intros [H_5 H_4].
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_din_bounded : forall x7, (0 <= x7) ->0 <= P_id_din x7.
          Proof.
            intros .
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_u32_bounded :
           forall x8 x10 x9 x7, 
            (0 <= x7) ->
             (0 <= x8) ->(0 <= x9) ->(0 <= x10) ->0 <= P_id_u32 x10 x9 x8 x7.
          Proof.
            intros .
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_dout_bounded : forall x7, (0 <= x7) ->0 <= P_id_dout x7.
          Proof.
            intros .
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_plus_bounded :
           forall x8 x7, (0 <= x7) ->(0 <= x8) ->0 <= P_id_plus x8 x7.
          Proof.
            intros .
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_u42_bounded :
           forall x8 x9 x7, 
            (0 <= x7) ->(0 <= x8) ->(0 <= x9) ->0 <= P_id_u42 x9 x8 x7.
          Proof.
            intros .
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_times_bounded :
           forall x8 x7, (0 <= x7) ->(0 <= x8) ->0 <= P_id_times x8 x7.
          Proof.
            intros .
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_der_bounded : forall x7, (0 <= x7) ->0 <= P_id_der x7.
          Proof.
            intros .
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_u41_bounded :
           forall x8 x7, (0 <= x7) ->(0 <= x8) ->0 <= P_id_u41 x8 x7.
          Proof.
            intros .
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_u22_bounded :
           forall x8 x10 x9 x7, 
            (0 <= x7) ->
             (0 <= x8) ->(0 <= x9) ->(0 <= x10) ->0 <= P_id_u22 x10 x9 x8 x7.
          Proof.
            intros .
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_u21_bounded :
           forall x8 x9 x7, 
            (0 <= x7) ->(0 <= x8) ->(0 <= x9) ->0 <= P_id_u21 x9 x8 x7.
          Proof.
            intros .
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_u31_bounded :
           forall x8 x9 x7, 
            (0 <= x7) ->(0 <= x8) ->(0 <= x9) ->0 <= P_id_u31 x9 x8 x7.
          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_din P_id_u32 P_id_dout P_id_plus P_id_u42 
             P_id_times P_id_der P_id_u41 P_id_u22 P_id_u21 P_id_u31.
          
          Lemma measure_equation :
           forall t, 
            measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_din (x7::nil)) =>
                           P_id_din (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u32 (x10::x9::x8::
                             x7::nil)) =>
                           P_id_u32 (measure x10) (measure x9) (measure x8) 
                            (measure x7)
                          | (algebra.Alg.Term algebra.F.id_dout (x7::nil)) =>
                           P_id_dout (measure x7)
                          | (algebra.Alg.Term algebra.F.id_plus (x8::
                             x7::nil)) =>
                           P_id_plus (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u42 (x9::x8::
                             x7::nil)) =>
                           P_id_u42 (measure x9) (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_times (x8::
                             x7::nil)) =>
                           P_id_times (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_der (x7::nil)) =>
                           P_id_der (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil)) =>
                           P_id_u41 (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u22 (x10::x9::x8::
                             x7::nil)) =>
                           P_id_u22 (measure x10) (measure x9) (measure x8) 
                            (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u21 (x9::x8::
                             x7::nil)) =>
                           P_id_u21 (measure x9) (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u31 (x9::x8::
                             x7::nil)) =>
                           P_id_u31 (measure x9) (measure x8) (measure x7)
                          | _ => 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_din_monotonic;assumption.
            intros ;apply P_id_u32_monotonic;assumption.
            intros ;apply P_id_dout_monotonic;assumption.
            intros ;apply P_id_plus_monotonic;assumption.
            intros ;apply P_id_u42_monotonic;assumption.
            intros ;apply P_id_times_monotonic;assumption.
            intros ;apply P_id_der_monotonic;assumption.
            intros ;apply P_id_u41_monotonic;assumption.
            intros ;apply P_id_u22_monotonic;assumption.
            intros ;apply P_id_u21_monotonic;assumption.
            intros ;apply P_id_u31_monotonic;assumption.
            intros ;apply P_id_din_bounded;assumption.
            intros ;apply P_id_u32_bounded;assumption.
            intros ;apply P_id_dout_bounded;assumption.
            intros ;apply P_id_plus_bounded;assumption.
            intros ;apply P_id_u42_bounded;assumption.
            intros ;apply P_id_times_bounded;assumption.
            intros ;apply P_id_der_bounded;assumption.
            intros ;apply P_id_u41_bounded;assumption.
            intros ;apply P_id_u22_bounded;assumption.
            intros ;apply P_id_u21_bounded;assumption.
            intros ;apply P_id_u31_bounded;assumption.
            apply rules_monotonic.
          Qed.
          
          Definition P_id_U41 (x7:Z) (x8:Z) := 0.
          
          Definition P_id_U21 (x7:Z) (x8:Z) (x9:Z) := 0.
          
          Definition P_id_U31 (x7:Z) (x8:Z) (x9:Z) := 0.
          
          Definition P_id_U42 (x7:Z) (x8:Z) (x9:Z) := 0.
          
          Definition P_id_U22 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 0.
          
          Definition P_id_DIN (x7:Z) := 1* x7.
          
          Definition P_id_U32 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 0.
          
          Lemma P_id_U41_monotonic :
           forall x8 x10 x9 x7, 
            (0 <= x10)/\ (x10 <= x9) ->
             (0 <= x8)/\ (x8 <= x7) ->P_id_U41 x8 x10 <= P_id_U41 x7 x9.
          Proof.
            intros x10 x9 x8 x7.
            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_U21_monotonic :
           forall x8 x12 x10 x9 x11 x7, 
            (0 <= x12)/\ (x12 <= x11) ->
             (0 <= x10)/\ (x10 <= x9) ->
              (0 <= x8)/\ (x8 <= x7) ->
               P_id_U21 x8 x10 x12 <= P_id_U21 x7 x9 x11.
          Proof.
            intros x12 x11 x10 x9 x8 x7.
            intros [H_1 H_0].
            intros [H_3 H_2].
            intros [H_5 H_4].
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_U31_monotonic :
           forall x8 x12 x10 x9 x11 x7, 
            (0 <= x12)/\ (x12 <= x11) ->
             (0 <= x10)/\ (x10 <= x9) ->
              (0 <= x8)/\ (x8 <= x7) ->
               P_id_U31 x8 x10 x12 <= P_id_U31 x7 x9 x11.
          Proof.
            intros x12 x11 x10 x9 x8 x7.
            intros [H_1 H_0].
            intros [H_3 H_2].
            intros [H_5 H_4].
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_U42_monotonic :
           forall x8 x12 x10 x9 x11 x7, 
            (0 <= x12)/\ (x12 <= x11) ->
             (0 <= x10)/\ (x10 <= x9) ->
              (0 <= x8)/\ (x8 <= x7) ->
               P_id_U42 x8 x10 x12 <= P_id_U42 x7 x9 x11.
          Proof.
            intros x12 x11 x10 x9 x8 x7.
            intros [H_1 H_0].
            intros [H_3 H_2].
            intros [H_5 H_4].
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_U22_monotonic :
           forall x8 x12 x10 x14 x9 x13 x11 x7, 
            (0 <= x14)/\ (x14 <= x13) ->
             (0 <= x12)/\ (x12 <= x11) ->
              (0 <= x10)/\ (x10 <= x9) ->
               (0 <= x8)/\ (x8 <= x7) ->
                P_id_U22 x8 x10 x12 x14 <= P_id_U22 x7 x9 x11 x13.
          Proof.
            intros x14 x13 x12 x11 x10 x9 x8 x7.
            intros [H_1 H_0].
            intros [H_3 H_2].
            intros [H_5 H_4].
            intros [H_7 H_6].
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_DIN_monotonic :
           forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_DIN x8 <= P_id_DIN x7.
          Proof.
            intros x8 x7.
            intros [H_1 H_0].
            
            cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
             (auto with zarith)||(repeat (translate_vars );prove_ineq ).
          Qed.
          
          Lemma P_id_U32_monotonic :
           forall x8 x12 x10 x14 x9 x13 x11 x7, 
            (0 <= x14)/\ (x14 <= x13) ->
             (0 <= x12)/\ (x12 <= x11) ->
              (0 <= x10)/\ (x10 <= x9) ->
               (0 <= x8)/\ (x8 <= x7) ->
                P_id_U32 x8 x10 x12 x14 <= P_id_U32 x7 x9 x11 x13.
          Proof.
            intros x14 x13 x12 x11 x10 x9 x8 x7.
            intros [H_1 H_0].
            intros [H_3 H_2].
            intros [H_5 H_4].
            intros [H_7 H_6].
            
            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_din P_id_u32 P_id_dout P_id_plus 
             P_id_u42 P_id_times P_id_der P_id_u41 P_id_u22 P_id_u21 
             P_id_u31 P_id_U41 P_id_U21 P_id_U31 P_id_U42 P_id_U22 P_id_DIN 
             P_id_U32.
          
          Lemma marked_measure_equation :
           forall t, 
            marked_measure t = match t with
                                 | (algebra.Alg.Term algebra.F.id_u41 (x8::
                                    x7::nil)) =>
                                  P_id_U41 (measure x8) (measure x7)
                                 | (algebra.Alg.Term algebra.F.id_u21 (x9::
                                    x8::x7::nil)) =>
                                  P_id_U21 (measure x9) (measure x8) 
                                   (measure x7)
                                 | (algebra.Alg.Term algebra.F.id_u31 (x9::
                                    x8::x7::nil)) =>
                                  P_id_U31 (measure x9) (measure x8) 
                                   (measure x7)
                                 | (algebra.Alg.Term algebra.F.id_u42 (x9::
                                    x8::x7::nil)) =>
                                  P_id_U42 (measure x9) (measure x8) 
                                   (measure x7)
                                 | (algebra.Alg.Term algebra.F.id_u22 (x10::
                                    x9::x8::x7::nil)) =>
                                  P_id_U22 (measure x10) (measure x9) 
                                   (measure x8) (measure x7)
                                 | (algebra.Alg.Term algebra.F.id_din 
                                    (x7::nil)) =>
                                  P_id_DIN (measure x7)
                                 | (algebra.Alg.Term algebra.F.id_u32 (x10::
                                    x9::x8::x7::nil)) =>
                                  P_id_U32 (measure x10) (measure x9) 
                                   (measure x8) (measure x7)
                                 | _ => 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_din_monotonic;assumption.
            intros ;apply P_id_u32_monotonic;assumption.
            intros ;apply P_id_dout_monotonic;assumption.
            intros ;apply P_id_plus_monotonic;assumption.
            intros ;apply P_id_u42_monotonic;assumption.
            intros ;apply P_id_times_monotonic;assumption.
            intros ;apply P_id_der_monotonic;assumption.
            intros ;apply P_id_u41_monotonic;assumption.
            intros ;apply P_id_u22_monotonic;assumption.
            intros ;apply P_id_u21_monotonic;assumption.
            intros ;apply P_id_u31_monotonic;assumption.
            intros ;apply P_id_din_bounded;assumption.
            intros ;apply P_id_u32_bounded;assumption.
            intros ;apply P_id_dout_bounded;assumption.
            intros ;apply P_id_plus_bounded;assumption.
            intros ;apply P_id_u42_bounded;assumption.
            intros ;apply P_id_times_bounded;assumption.
            intros ;apply P_id_der_bounded;assumption.
            intros ;apply P_id_u41_bounded;assumption.
            intros ;apply P_id_u22_bounded;assumption.
            intros ;apply P_id_u21_bounded;assumption.
            intros ;apply P_id_u31_bounded;assumption.
            apply rules_monotonic.
            intros ;apply P_id_U41_monotonic;assumption.
            intros ;apply P_id_U21_monotonic;assumption.
            intros ;apply P_id_U31_monotonic;assumption.
            intros ;apply P_id_U42_monotonic;assumption.
            intros ;apply P_id_U22_monotonic;assumption.
            intros ;apply P_id_DIN_monotonic;assumption.
            intros ;apply P_id_U32_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_scc_4_large_scc_2_large_scc_2_large_scc_2.DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_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_scc_4_large_scc_2_large_scc_2_large_scc_2_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_din (x7:Z) := 2.
         
         Definition P_id_u32 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 0.
         
         Definition P_id_dout (x7:Z) := 0.
         
         Definition P_id_plus (x7:Z) (x8:Z) := 3 + 1* x7.
         
         Definition P_id_u42 (x7:Z) (x8:Z) (x9:Z) := 0.
         
         Definition P_id_times (x7:Z) (x8:Z) := 3 + 1* x7.
         
         Definition P_id_der (x7:Z) := 2* x7.
         
         Definition P_id_u41 (x7:Z) (x8:Z) := 0.
         
         Definition P_id_u22 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 0.
         
         Definition P_id_u21 (x7:Z) (x8:Z) (x9:Z) := 0.
         
         Definition P_id_u31 (x7:Z) (x8:Z) (x9:Z) := 1.
         
         Lemma P_id_din_monotonic :
          forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_din x8 <= P_id_din x7.
         Proof.
           intros x8 x7.
           intros [H_1 H_0].
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_u32_monotonic :
          forall x8 x12 x10 x14 x9 x13 x11 x7, 
           (0 <= x14)/\ (x14 <= x13) ->
            (0 <= x12)/\ (x12 <= x11) ->
             (0 <= x10)/\ (x10 <= x9) ->
              (0 <= x8)/\ (x8 <= x7) ->
               P_id_u32 x8 x10 x12 x14 <= P_id_u32 x7 x9 x11 x13.
         Proof.
           intros x14 x13 x12 x11 x10 x9 x8 x7.
           intros [H_1 H_0].
           intros [H_3 H_2].
           intros [H_5 H_4].
           intros [H_7 H_6].
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_dout_monotonic :
          forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_dout x8 <= P_id_dout x7.
         Proof.
           intros x8 x7.
           intros [H_1 H_0].
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_plus_monotonic :
          forall x8 x10 x9 x7, 
           (0 <= x10)/\ (x10 <= x9) ->
            (0 <= x8)/\ (x8 <= x7) ->P_id_plus x8 x10 <= P_id_plus x7 x9.
         Proof.
           intros x10 x9 x8 x7.
           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_u42_monotonic :
          forall x8 x12 x10 x9 x11 x7, 
           (0 <= x12)/\ (x12 <= x11) ->
            (0 <= x10)/\ (x10 <= x9) ->
             (0 <= x8)/\ (x8 <= x7) ->
              P_id_u42 x8 x10 x12 <= P_id_u42 x7 x9 x11.
         Proof.
           intros x12 x11 x10 x9 x8 x7.
           intros [H_1 H_0].
           intros [H_3 H_2].
           intros [H_5 H_4].
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_times_monotonic :
          forall x8 x10 x9 x7, 
           (0 <= x10)/\ (x10 <= x9) ->
            (0 <= x8)/\ (x8 <= x7) ->P_id_times x8 x10 <= P_id_times x7 x9.
         Proof.
           intros x10 x9 x8 x7.
           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_der_monotonic :
          forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_der x8 <= P_id_der x7.
         Proof.
           intros x8 x7.
           intros [H_1 H_0].
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_u41_monotonic :
          forall x8 x10 x9 x7, 
           (0 <= x10)/\ (x10 <= x9) ->
            (0 <= x8)/\ (x8 <= x7) ->P_id_u41 x8 x10 <= P_id_u41 x7 x9.
         Proof.
           intros x10 x9 x8 x7.
           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_u22_monotonic :
          forall x8 x12 x10 x14 x9 x13 x11 x7, 
           (0 <= x14)/\ (x14 <= x13) ->
            (0 <= x12)/\ (x12 <= x11) ->
             (0 <= x10)/\ (x10 <= x9) ->
              (0 <= x8)/\ (x8 <= x7) ->
               P_id_u22 x8 x10 x12 x14 <= P_id_u22 x7 x9 x11 x13.
         Proof.
           intros x14 x13 x12 x11 x10 x9 x8 x7.
           intros [H_1 H_0].
           intros [H_3 H_2].
           intros [H_5 H_4].
           intros [H_7 H_6].
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_u21_monotonic :
          forall x8 x12 x10 x9 x11 x7, 
           (0 <= x12)/\ (x12 <= x11) ->
            (0 <= x10)/\ (x10 <= x9) ->
             (0 <= x8)/\ (x8 <= x7) ->
              P_id_u21 x8 x10 x12 <= P_id_u21 x7 x9 x11.
         Proof.
           intros x12 x11 x10 x9 x8 x7.
           intros [H_1 H_0].
           intros [H_3 H_2].
           intros [H_5 H_4].
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_u31_monotonic :
          forall x8 x12 x10 x9 x11 x7, 
           (0 <= x12)/\ (x12 <= x11) ->
            (0 <= x10)/\ (x10 <= x9) ->
             (0 <= x8)/\ (x8 <= x7) ->
              P_id_u31 x8 x10 x12 <= P_id_u31 x7 x9 x11.
         Proof.
           intros x12 x11 x10 x9 x8 x7.
           intros [H_1 H_0].
           intros [H_3 H_2].
           intros [H_5 H_4].
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_din_bounded : forall x7, (0 <= x7) ->0 <= P_id_din x7.
         Proof.
           intros .
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_u32_bounded :
          forall x8 x10 x9 x7, 
           (0 <= x7) ->
            (0 <= x8) ->(0 <= x9) ->(0 <= x10) ->0 <= P_id_u32 x10 x9 x8 x7.
         Proof.
           intros .
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_dout_bounded : forall x7, (0 <= x7) ->0 <= P_id_dout x7.
         Proof.
           intros .
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_plus_bounded :
          forall x8 x7, (0 <= x7) ->(0 <= x8) ->0 <= P_id_plus x8 x7.
         Proof.
           intros .
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_u42_bounded :
          forall x8 x9 x7, 
           (0 <= x7) ->(0 <= x8) ->(0 <= x9) ->0 <= P_id_u42 x9 x8 x7.
         Proof.
           intros .
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_times_bounded :
          forall x8 x7, (0 <= x7) ->(0 <= x8) ->0 <= P_id_times x8 x7.
         Proof.
           intros .
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_der_bounded : forall x7, (0 <= x7) ->0 <= P_id_der x7.
         Proof.
           intros .
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_u41_bounded :
          forall x8 x7, (0 <= x7) ->(0 <= x8) ->0 <= P_id_u41 x8 x7.
         Proof.
           intros .
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_u22_bounded :
          forall x8 x10 x9 x7, 
           (0 <= x7) ->
            (0 <= x8) ->(0 <= x9) ->(0 <= x10) ->0 <= P_id_u22 x10 x9 x8 x7.
         Proof.
           intros .
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_u21_bounded :
          forall x8 x9 x7, 
           (0 <= x7) ->(0 <= x8) ->(0 <= x9) ->0 <= P_id_u21 x9 x8 x7.
         Proof.
           intros .
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_u31_bounded :
          forall x8 x9 x7, 
           (0 <= x7) ->(0 <= x8) ->(0 <= x9) ->0 <= P_id_u31 x9 x8 x7.
         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_din P_id_u32 P_id_dout P_id_plus P_id_u42 
            P_id_times P_id_der P_id_u41 P_id_u22 P_id_u21 P_id_u31.
         
         Lemma measure_equation :
          forall t, 
           measure t = match t with
                         | (algebra.Alg.Term algebra.F.id_din (x7::nil)) =>
                          P_id_din (measure x7)
                         | (algebra.Alg.Term algebra.F.id_u32 (x10::x9::x8::
                            x7::nil)) =>
                          P_id_u32 (measure x10) (measure x9) (measure x8) 
                           (measure x7)
                         | (algebra.Alg.Term algebra.F.id_dout (x7::nil)) =>
                          P_id_dout (measure x7)
                         | (algebra.Alg.Term algebra.F.id_plus (x8::x7::nil)) =>
                          P_id_plus (measure x8) (measure x7)
                         | (algebra.Alg.Term algebra.F.id_u42 (x9::x8::
                            x7::nil)) =>
                          P_id_u42 (measure x9) (measure x8) (measure x7)
                         | (algebra.Alg.Term algebra.F.id_times (x8::
                            x7::nil)) =>
                          P_id_times (measure x8) (measure x7)
                         | (algebra.Alg.Term algebra.F.id_der (x7::nil)) =>
                          P_id_der (measure x7)
                         | (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil)) =>
                          P_id_u41 (measure x8) (measure x7)
                         | (algebra.Alg.Term algebra.F.id_u22 (x10::x9::x8::
                            x7::nil)) =>
                          P_id_u22 (measure x10) (measure x9) (measure x8) 
                           (measure x7)
                         | (algebra.Alg.Term algebra.F.id_u21 (x9::x8::
                            x7::nil)) =>
                          P_id_u21 (measure x9) (measure x8) (measure x7)
                         | (algebra.Alg.Term algebra.F.id_u31 (x9::x8::
                            x7::nil)) =>
                          P_id_u31 (measure x9) (measure x8) (measure x7)
                         | _ => 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_din_monotonic;assumption.
           intros ;apply P_id_u32_monotonic;assumption.
           intros ;apply P_id_dout_monotonic;assumption.
           intros ;apply P_id_plus_monotonic;assumption.
           intros ;apply P_id_u42_monotonic;assumption.
           intros ;apply P_id_times_monotonic;assumption.
           intros ;apply P_id_der_monotonic;assumption.
           intros ;apply P_id_u41_monotonic;assumption.
           intros ;apply P_id_u22_monotonic;assumption.
           intros ;apply P_id_u21_monotonic;assumption.
           intros ;apply P_id_u31_monotonic;assumption.
           intros ;apply P_id_din_bounded;assumption.
           intros ;apply P_id_u32_bounded;assumption.
           intros ;apply P_id_dout_bounded;assumption.
           intros ;apply P_id_plus_bounded;assumption.
           intros ;apply P_id_u42_bounded;assumption.
           intros ;apply P_id_times_bounded;assumption.
           intros ;apply P_id_der_bounded;assumption.
           intros ;apply P_id_u41_bounded;assumption.
           intros ;apply P_id_u22_bounded;assumption.
           intros ;apply P_id_u21_bounded;assumption.
           intros ;apply P_id_u31_bounded;assumption.
           apply rules_monotonic.
         Qed.
         
         Definition P_id_U41 (x7:Z) (x8:Z) := 0.
         
         Definition P_id_U21 (x7:Z) (x8:Z) (x9:Z) := 0.
         
         Definition P_id_U31 (x7:Z) (x8:Z) (x9:Z) := 0.
         
         Definition P_id_U42 (x7:Z) (x8:Z) (x9:Z) := 0.
         
         Definition P_id_U22 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 0.
         
         Definition P_id_DIN (x7:Z) := 2* x7.
         
         Definition P_id_U32 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 0.
         
         Lemma P_id_U41_monotonic :
          forall x8 x10 x9 x7, 
           (0 <= x10)/\ (x10 <= x9) ->
            (0 <= x8)/\ (x8 <= x7) ->P_id_U41 x8 x10 <= P_id_U41 x7 x9.
         Proof.
           intros x10 x9 x8 x7.
           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_U21_monotonic :
          forall x8 x12 x10 x9 x11 x7, 
           (0 <= x12)/\ (x12 <= x11) ->
            (0 <= x10)/\ (x10 <= x9) ->
             (0 <= x8)/\ (x8 <= x7) ->
              P_id_U21 x8 x10 x12 <= P_id_U21 x7 x9 x11.
         Proof.
           intros x12 x11 x10 x9 x8 x7.
           intros [H_1 H_0].
           intros [H_3 H_2].
           intros [H_5 H_4].
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_U31_monotonic :
          forall x8 x12 x10 x9 x11 x7, 
           (0 <= x12)/\ (x12 <= x11) ->
            (0 <= x10)/\ (x10 <= x9) ->
             (0 <= x8)/\ (x8 <= x7) ->
              P_id_U31 x8 x10 x12 <= P_id_U31 x7 x9 x11.
         Proof.
           intros x12 x11 x10 x9 x8 x7.
           intros [H_1 H_0].
           intros [H_3 H_2].
           intros [H_5 H_4].
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_U42_monotonic :
          forall x8 x12 x10 x9 x11 x7, 
           (0 <= x12)/\ (x12 <= x11) ->
            (0 <= x10)/\ (x10 <= x9) ->
             (0 <= x8)/\ (x8 <= x7) ->
              P_id_U42 x8 x10 x12 <= P_id_U42 x7 x9 x11.
         Proof.
           intros x12 x11 x10 x9 x8 x7.
           intros [H_1 H_0].
           intros [H_3 H_2].
           intros [H_5 H_4].
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_U22_monotonic :
          forall x8 x12 x10 x14 x9 x13 x11 x7, 
           (0 <= x14)/\ (x14 <= x13) ->
            (0 <= x12)/\ (x12 <= x11) ->
             (0 <= x10)/\ (x10 <= x9) ->
              (0 <= x8)/\ (x8 <= x7) ->
               P_id_U22 x8 x10 x12 x14 <= P_id_U22 x7 x9 x11 x13.
         Proof.
           intros x14 x13 x12 x11 x10 x9 x8 x7.
           intros [H_1 H_0].
           intros [H_3 H_2].
           intros [H_5 H_4].
           intros [H_7 H_6].
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_DIN_monotonic :
          forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_DIN x8 <= P_id_DIN x7.
         Proof.
           intros x8 x7.
           intros [H_1 H_0].
           
           cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
            (auto with zarith)||(repeat (translate_vars );prove_ineq ).
         Qed.
         
         Lemma P_id_U32_monotonic :
          forall x8 x12 x10 x14 x9 x13 x11 x7, 
           (0 <= x14)/\ (x14 <= x13) ->
            (0 <= x12)/\ (x12 <= x11) ->
             (0 <= x10)/\ (x10 <= x9) ->
              (0 <= x8)/\ (x8 <= x7) ->
               P_id_U32 x8 x10 x12 x14 <= P_id_U32 x7 x9 x11 x13.
         Proof.
           intros x14 x13 x12 x11 x10 x9 x8 x7.
           intros [H_1 H_0].
           intros [H_3 H_2].
           intros [H_5 H_4].
           intros [H_7 H_6].
           
           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_din P_id_u32 P_id_dout P_id_plus 
            P_id_u42 P_id_times P_id_der P_id_u41 P_id_u22 P_id_u21 P_id_u31 
            P_id_U41 P_id_U21 P_id_U31 P_id_U42 P_id_U22 P_id_DIN P_id_U32.
         
         Lemma marked_measure_equation :
          forall t, 
           marked_measure t = match t with
                                | (algebra.Alg.Term algebra.F.id_u41 (x8::
                                   x7::nil)) =>
                                 P_id_U41 (measure x8) (measure x7)
                                | (algebra.Alg.Term algebra.F.id_u21 (x9::
                                   x8::x7::nil)) =>
                                 P_id_U21 (measure x9) (measure x8) 
                                  (measure x7)
                                | (algebra.Alg.Term algebra.F.id_u31 (x9::
                                   x8::x7::nil)) =>
                                 P_id_U31 (measure x9) (measure x8) 
                                  (measure x7)
                                | (algebra.Alg.Term algebra.F.id_u42 (x9::
                                   x8::x7::nil)) =>
                                 P_id_U42 (measure x9) (measure x8) 
                                  (measure x7)
                                | (algebra.Alg.Term algebra.F.id_u22 (x10::
                                   x9::x8::x7::nil)) =>
                                 P_id_U22 (measure x10) (measure x9) 
                                  (measure x8) (measure x7)
                                | (algebra.Alg.Term algebra.F.id_din 
                                   (x7::nil)) =>
                                 P_id_DIN (measure x7)
                                | (algebra.Alg.Term algebra.F.id_u32 (x10::
                                   x9::x8::x7::nil)) =>
                                 P_id_U32 (measure x10) (measure x9) 
                                  (measure x8) (measure x7)
                                | _ => 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_din_monotonic;assumption.
           intros ;apply P_id_u32_monotonic;assumption.
           intros ;apply P_id_dout_monotonic;assumption.
           intros ;apply P_id_plus_monotonic;assumption.
           intros ;apply P_id_u42_monotonic;assumption.
           intros ;apply P_id_times_monotonic;assumption.
           intros ;apply P_id_der_monotonic;assumption.
           intros ;apply P_id_u41_monotonic;assumption.
           intros ;apply P_id_u22_monotonic;assumption.
           intros ;apply P_id_u21_monotonic;assumption.
           intros ;apply P_id_u31_monotonic;assumption.
           intros ;apply P_id_din_bounded;assumption.
           intros ;apply P_id_u32_bounded;assumption.
           intros ;apply P_id_dout_bounded;assumption.
           intros ;apply P_id_plus_bounded;assumption.
           intros ;apply P_id_u42_bounded;assumption.
           intros ;apply P_id_times_bounded;assumption.
           intros ;apply P_id_der_bounded;assumption.
           intros ;apply P_id_u41_bounded;assumption.
           intros ;apply P_id_u22_bounded;assumption.
           intros ;apply P_id_u21_bounded;assumption.
           intros ;apply P_id_u31_bounded;assumption.
           apply rules_monotonic.
           intros ;apply P_id_U41_monotonic;assumption.
           intros ;apply P_id_U21_monotonic;assumption.
           intros ;apply P_id_U31_monotonic;assumption.
           intros ;apply P_id_U42_monotonic;assumption.
           intros ;apply P_id_U22_monotonic;assumption.
           intros ;apply P_id_DIN_monotonic;assumption.
           intros ;apply P_id_U32_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_scc_4_large_scc_2_large_scc_2_large_scc_2_strict_in_lt :
          Relation_Definitions.inclusion _ 
           DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_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_scc_4_large_scc_2_large_scc_2_large_scc_2_large_in_le :
          Relation_Definitions.inclusion _ 
           DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_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_scc_4_large_scc_2_large_scc_2_large_scc_2_large  := 
           WF_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_large.wf.
         
         
         Lemma wf :
          well_founded WF_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large.DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2
           .
         Proof.
           intros x.
           apply (well_founded_ind wf_lt).
           clear x.
           intros x.
           pattern x.
           
           apply (@Acc_ind _ 
                   DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_large).
           clear x.
           intros x _ IHx IHx'.
           constructor.
           intros y H.
           
           destruct H;
            (apply IHx';
              apply DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_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_scc_4_large_scc_2_large_scc_2_large_scc_2_large_in_le;
               econstructor eassumption])).
           
           apply wf_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2_large.
         Qed.
        End WF_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2.
        
        Definition wf_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2  := 
          WF_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2.wf.
        
        
        Lemma acc_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2 :
         forall x y, 
          (DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2 x y) ->
           Acc WF_DP_R_xml_0_scc_4_large_scc_2_large_scc_2.DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large 
            x.
        Proof.
          intros x.
          pattern x.
          
          apply (@Acc_ind _ 
                  DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2).
          intros x' _ Hrec y h.
          
          inversion h;clear h;subst;
           constructor;intros _y _h;inversion _h;clear _h;subst;
            (eapply Hrec;econstructor eassumption)||
            ((eapply acc_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_non_scc_1;
               econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
             ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||
              (eapply Hrec;
                econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))).
          apply wf_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2.
        Qed.
        
        
        Lemma wf :
         well_founded WF_DP_R_xml_0_scc_4_large_scc_2_large_scc_2.DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large
          .
        Proof.
          constructor;intros _y _h;inversion _h;clear _h;subst;
           (eapply acc_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_non_scc_1;
             econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
           ((eapply acc_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_non_scc_0;
              econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
            ((eapply acc_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_2;
               econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
             ((eapply acc_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_1;
                econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
              ((eapply acc_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large_scc_0;
                 econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
               ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||(fail)))))).
        Qed.
       End WF_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_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_din (x7:Z) := 0.
       
       Definition P_id_u32 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 3* x7.
       
       Definition P_id_dout (x7:Z) := 3.
       
       Definition P_id_plus (x7:Z) (x8:Z) := 0.
       
       Definition P_id_u42 (x7:Z) (x8:Z) (x9:Z) := 1* x7.
       
       Definition P_id_times (x7:Z) (x8:Z) := 0.
       
       Definition P_id_der (x7:Z) := 0.
       
       Definition P_id_u41 (x7:Z) (x8:Z) := 0.
       
       Definition P_id_u22 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 2 + 2* x7.
       
       Definition P_id_u21 (x7:Z) (x8:Z) (x9:Z) := 3* x7.
       
       Definition P_id_u31 (x7:Z) (x8:Z) (x9:Z) := 0.
       
       Lemma P_id_din_monotonic :
        forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_din x8 <= P_id_din x7.
       Proof.
         intros x8 x7.
         intros [H_1 H_0].
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_u32_monotonic :
        forall x8 x12 x10 x14 x9 x13 x11 x7, 
         (0 <= x14)/\ (x14 <= x13) ->
          (0 <= x12)/\ (x12 <= x11) ->
           (0 <= x10)/\ (x10 <= x9) ->
            (0 <= x8)/\ (x8 <= x7) ->
             P_id_u32 x8 x10 x12 x14 <= P_id_u32 x7 x9 x11 x13.
       Proof.
         intros x14 x13 x12 x11 x10 x9 x8 x7.
         intros [H_1 H_0].
         intros [H_3 H_2].
         intros [H_5 H_4].
         intros [H_7 H_6].
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_dout_monotonic :
        forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_dout x8 <= P_id_dout x7.
       Proof.
         intros x8 x7.
         intros [H_1 H_0].
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_plus_monotonic :
        forall x8 x10 x9 x7, 
         (0 <= x10)/\ (x10 <= x9) ->
          (0 <= x8)/\ (x8 <= x7) ->P_id_plus x8 x10 <= P_id_plus x7 x9.
       Proof.
         intros x10 x9 x8 x7.
         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_u42_monotonic :
        forall x8 x12 x10 x9 x11 x7, 
         (0 <= x12)/\ (x12 <= x11) ->
          (0 <= x10)/\ (x10 <= x9) ->
           (0 <= x8)/\ (x8 <= x7) ->P_id_u42 x8 x10 x12 <= P_id_u42 x7 x9 x11.
       Proof.
         intros x12 x11 x10 x9 x8 x7.
         intros [H_1 H_0].
         intros [H_3 H_2].
         intros [H_5 H_4].
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_times_monotonic :
        forall x8 x10 x9 x7, 
         (0 <= x10)/\ (x10 <= x9) ->
          (0 <= x8)/\ (x8 <= x7) ->P_id_times x8 x10 <= P_id_times x7 x9.
       Proof.
         intros x10 x9 x8 x7.
         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_der_monotonic :
        forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_der x8 <= P_id_der x7.
       Proof.
         intros x8 x7.
         intros [H_1 H_0].
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_u41_monotonic :
        forall x8 x10 x9 x7, 
         (0 <= x10)/\ (x10 <= x9) ->
          (0 <= x8)/\ (x8 <= x7) ->P_id_u41 x8 x10 <= P_id_u41 x7 x9.
       Proof.
         intros x10 x9 x8 x7.
         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_u22_monotonic :
        forall x8 x12 x10 x14 x9 x13 x11 x7, 
         (0 <= x14)/\ (x14 <= x13) ->
          (0 <= x12)/\ (x12 <= x11) ->
           (0 <= x10)/\ (x10 <= x9) ->
            (0 <= x8)/\ (x8 <= x7) ->
             P_id_u22 x8 x10 x12 x14 <= P_id_u22 x7 x9 x11 x13.
       Proof.
         intros x14 x13 x12 x11 x10 x9 x8 x7.
         intros [H_1 H_0].
         intros [H_3 H_2].
         intros [H_5 H_4].
         intros [H_7 H_6].
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_u21_monotonic :
        forall x8 x12 x10 x9 x11 x7, 
         (0 <= x12)/\ (x12 <= x11) ->
          (0 <= x10)/\ (x10 <= x9) ->
           (0 <= x8)/\ (x8 <= x7) ->P_id_u21 x8 x10 x12 <= P_id_u21 x7 x9 x11.
       Proof.
         intros x12 x11 x10 x9 x8 x7.
         intros [H_1 H_0].
         intros [H_3 H_2].
         intros [H_5 H_4].
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_u31_monotonic :
        forall x8 x12 x10 x9 x11 x7, 
         (0 <= x12)/\ (x12 <= x11) ->
          (0 <= x10)/\ (x10 <= x9) ->
           (0 <= x8)/\ (x8 <= x7) ->P_id_u31 x8 x10 x12 <= P_id_u31 x7 x9 x11.
       Proof.
         intros x12 x11 x10 x9 x8 x7.
         intros [H_1 H_0].
         intros [H_3 H_2].
         intros [H_5 H_4].
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_din_bounded : forall x7, (0 <= x7) ->0 <= P_id_din x7.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_u32_bounded :
        forall x8 x10 x9 x7, 
         (0 <= x7) ->
          (0 <= x8) ->(0 <= x9) ->(0 <= x10) ->0 <= P_id_u32 x10 x9 x8 x7.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_dout_bounded : forall x7, (0 <= x7) ->0 <= P_id_dout x7.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_plus_bounded :
        forall x8 x7, (0 <= x7) ->(0 <= x8) ->0 <= P_id_plus x8 x7.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_u42_bounded :
        forall x8 x9 x7, 
         (0 <= x7) ->(0 <= x8) ->(0 <= x9) ->0 <= P_id_u42 x9 x8 x7.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_times_bounded :
        forall x8 x7, (0 <= x7) ->(0 <= x8) ->0 <= P_id_times x8 x7.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_der_bounded : forall x7, (0 <= x7) ->0 <= P_id_der x7.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_u41_bounded :
        forall x8 x7, (0 <= x7) ->(0 <= x8) ->0 <= P_id_u41 x8 x7.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_u22_bounded :
        forall x8 x10 x9 x7, 
         (0 <= x7) ->
          (0 <= x8) ->(0 <= x9) ->(0 <= x10) ->0 <= P_id_u22 x10 x9 x8 x7.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_u21_bounded :
        forall x8 x9 x7, 
         (0 <= x7) ->(0 <= x8) ->(0 <= x9) ->0 <= P_id_u21 x9 x8 x7.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_u31_bounded :
        forall x8 x9 x7, 
         (0 <= x7) ->(0 <= x8) ->(0 <= x9) ->0 <= P_id_u31 x9 x8 x7.
       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_din P_id_u32 P_id_dout P_id_plus P_id_u42 
          P_id_times P_id_der P_id_u41 P_id_u22 P_id_u21 P_id_u31.
       
       Lemma measure_equation :
        forall t, 
         measure t = match t with
                       | (algebra.Alg.Term algebra.F.id_din (x7::nil)) =>
                        P_id_din (measure x7)
                       | (algebra.Alg.Term algebra.F.id_u32 (x10::x9::x8::
                          x7::nil)) =>
                        P_id_u32 (measure x10) (measure x9) (measure x8) 
                         (measure x7)
                       | (algebra.Alg.Term algebra.F.id_dout (x7::nil)) =>
                        P_id_dout (measure x7)
                       | (algebra.Alg.Term algebra.F.id_plus (x8::x7::nil)) =>
                        P_id_plus (measure x8) (measure x7)
                       | (algebra.Alg.Term algebra.F.id_u42 (x9::x8::
                          x7::nil)) =>
                        P_id_u42 (measure x9) (measure x8) (measure x7)
                       | (algebra.Alg.Term algebra.F.id_times (x8::x7::nil)) =>
                        P_id_times (measure x8) (measure x7)
                       | (algebra.Alg.Term algebra.F.id_der (x7::nil)) =>
                        P_id_der (measure x7)
                       | (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil)) =>
                        P_id_u41 (measure x8) (measure x7)
                       | (algebra.Alg.Term algebra.F.id_u22 (x10::x9::x8::
                          x7::nil)) =>
                        P_id_u22 (measure x10) (measure x9) (measure x8) 
                         (measure x7)
                       | (algebra.Alg.Term algebra.F.id_u21 (x9::x8::
                          x7::nil)) =>
                        P_id_u21 (measure x9) (measure x8) (measure x7)
                       | (algebra.Alg.Term algebra.F.id_u31 (x9::x8::
                          x7::nil)) =>
                        P_id_u31 (measure x9) (measure x8) (measure x7)
                       | _ => 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_din_monotonic;assumption.
         intros ;apply P_id_u32_monotonic;assumption.
         intros ;apply P_id_dout_monotonic;assumption.
         intros ;apply P_id_plus_monotonic;assumption.
         intros ;apply P_id_u42_monotonic;assumption.
         intros ;apply P_id_times_monotonic;assumption.
         intros ;apply P_id_der_monotonic;assumption.
         intros ;apply P_id_u41_monotonic;assumption.
         intros ;apply P_id_u22_monotonic;assumption.
         intros ;apply P_id_u21_monotonic;assumption.
         intros ;apply P_id_u31_monotonic;assumption.
         intros ;apply P_id_din_bounded;assumption.
         intros ;apply P_id_u32_bounded;assumption.
         intros ;apply P_id_dout_bounded;assumption.
         intros ;apply P_id_plus_bounded;assumption.
         intros ;apply P_id_u42_bounded;assumption.
         intros ;apply P_id_times_bounded;assumption.
         intros ;apply P_id_der_bounded;assumption.
         intros ;apply P_id_u41_bounded;assumption.
         intros ;apply P_id_u22_bounded;assumption.
         intros ;apply P_id_u21_bounded;assumption.
         intros ;apply P_id_u31_bounded;assumption.
         apply rules_monotonic.
       Qed.
       
       Definition P_id_U41 (x7:Z) (x8:Z) := 2* x7.
       
       Definition P_id_U21 (x7:Z) (x8:Z) (x9:Z) := 0.
       
       Definition P_id_U31 (x7:Z) (x8:Z) (x9:Z) := 0.
       
       Definition P_id_U42 (x7:Z) (x8:Z) (x9:Z) := 0.
       
       Definition P_id_U22 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 0.
       
       Definition P_id_DIN (x7:Z) := 0.
       
       Definition P_id_U32 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 0.
       
       Lemma P_id_U41_monotonic :
        forall x8 x10 x9 x7, 
         (0 <= x10)/\ (x10 <= x9) ->
          (0 <= x8)/\ (x8 <= x7) ->P_id_U41 x8 x10 <= P_id_U41 x7 x9.
       Proof.
         intros x10 x9 x8 x7.
         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_U21_monotonic :
        forall x8 x12 x10 x9 x11 x7, 
         (0 <= x12)/\ (x12 <= x11) ->
          (0 <= x10)/\ (x10 <= x9) ->
           (0 <= x8)/\ (x8 <= x7) ->P_id_U21 x8 x10 x12 <= P_id_U21 x7 x9 x11.
       Proof.
         intros x12 x11 x10 x9 x8 x7.
         intros [H_1 H_0].
         intros [H_3 H_2].
         intros [H_5 H_4].
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_U31_monotonic :
        forall x8 x12 x10 x9 x11 x7, 
         (0 <= x12)/\ (x12 <= x11) ->
          (0 <= x10)/\ (x10 <= x9) ->
           (0 <= x8)/\ (x8 <= x7) ->P_id_U31 x8 x10 x12 <= P_id_U31 x7 x9 x11.
       Proof.
         intros x12 x11 x10 x9 x8 x7.
         intros [H_1 H_0].
         intros [H_3 H_2].
         intros [H_5 H_4].
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_U42_monotonic :
        forall x8 x12 x10 x9 x11 x7, 
         (0 <= x12)/\ (x12 <= x11) ->
          (0 <= x10)/\ (x10 <= x9) ->
           (0 <= x8)/\ (x8 <= x7) ->P_id_U42 x8 x10 x12 <= P_id_U42 x7 x9 x11.
       Proof.
         intros x12 x11 x10 x9 x8 x7.
         intros [H_1 H_0].
         intros [H_3 H_2].
         intros [H_5 H_4].
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_U22_monotonic :
        forall x8 x12 x10 x14 x9 x13 x11 x7, 
         (0 <= x14)/\ (x14 <= x13) ->
          (0 <= x12)/\ (x12 <= x11) ->
           (0 <= x10)/\ (x10 <= x9) ->
            (0 <= x8)/\ (x8 <= x7) ->
             P_id_U22 x8 x10 x12 x14 <= P_id_U22 x7 x9 x11 x13.
       Proof.
         intros x14 x13 x12 x11 x10 x9 x8 x7.
         intros [H_1 H_0].
         intros [H_3 H_2].
         intros [H_5 H_4].
         intros [H_7 H_6].
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_DIN_monotonic :
        forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_DIN x8 <= P_id_DIN x7.
       Proof.
         intros x8 x7.
         intros [H_1 H_0].
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_U32_monotonic :
        forall x8 x12 x10 x14 x9 x13 x11 x7, 
         (0 <= x14)/\ (x14 <= x13) ->
          (0 <= x12)/\ (x12 <= x11) ->
           (0 <= x10)/\ (x10 <= x9) ->
            (0 <= x8)/\ (x8 <= x7) ->
             P_id_U32 x8 x10 x12 x14 <= P_id_U32 x7 x9 x11 x13.
       Proof.
         intros x14 x13 x12 x11 x10 x9 x8 x7.
         intros [H_1 H_0].
         intros [H_3 H_2].
         intros [H_5 H_4].
         intros [H_7 H_6].
         
         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_din P_id_u32 P_id_dout P_id_plus 
          P_id_u42 P_id_times P_id_der P_id_u41 P_id_u22 P_id_u21 P_id_u31 
          P_id_U41 P_id_U21 P_id_U31 P_id_U42 P_id_U22 P_id_DIN P_id_U32.
       
       Lemma marked_measure_equation :
        forall t, 
         marked_measure t = match t with
                              | (algebra.Alg.Term algebra.F.id_u41 (x8::
                                 x7::nil)) =>
                               P_id_U41 (measure x8) (measure x7)
                              | (algebra.Alg.Term algebra.F.id_u21 (x9::x8::
                                 x7::nil)) =>
                               P_id_U21 (measure x9) (measure x8) 
                                (measure x7)
                              | (algebra.Alg.Term algebra.F.id_u31 (x9::x8::
                                 x7::nil)) =>
                               P_id_U31 (measure x9) (measure x8) 
                                (measure x7)
                              | (algebra.Alg.Term algebra.F.id_u42 (x9::x8::
                                 x7::nil)) =>
                               P_id_U42 (measure x9) (measure x8) 
                                (measure x7)
                              | (algebra.Alg.Term algebra.F.id_u22 (x10::x9::
                                 x8::x7::nil)) =>
                               P_id_U22 (measure x10) (measure x9) 
                                (measure x8) (measure x7)
                              | (algebra.Alg.Term algebra.F.id_din (x7::nil)) =>
                               P_id_DIN (measure x7)
                              | (algebra.Alg.Term algebra.F.id_u32 (x10::x9::
                                 x8::x7::nil)) =>
                               P_id_U32 (measure x10) (measure x9) 
                                (measure x8) (measure x7)
                              | _ => 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_din_monotonic;assumption.
         intros ;apply P_id_u32_monotonic;assumption.
         intros ;apply P_id_dout_monotonic;assumption.
         intros ;apply P_id_plus_monotonic;assumption.
         intros ;apply P_id_u42_monotonic;assumption.
         intros ;apply P_id_times_monotonic;assumption.
         intros ;apply P_id_der_monotonic;assumption.
         intros ;apply P_id_u41_monotonic;assumption.
         intros ;apply P_id_u22_monotonic;assumption.
         intros ;apply P_id_u21_monotonic;assumption.
         intros ;apply P_id_u31_monotonic;assumption.
         intros ;apply P_id_din_bounded;assumption.
         intros ;apply P_id_u32_bounded;assumption.
         intros ;apply P_id_dout_bounded;assumption.
         intros ;apply P_id_plus_bounded;assumption.
         intros ;apply P_id_u42_bounded;assumption.
         intros ;apply P_id_times_bounded;assumption.
         intros ;apply P_id_der_bounded;assumption.
         intros ;apply P_id_u41_bounded;assumption.
         intros ;apply P_id_u22_bounded;assumption.
         intros ;apply P_id_u21_bounded;assumption.
         intros ;apply P_id_u31_bounded;assumption.
         apply rules_monotonic.
         intros ;apply P_id_U41_monotonic;assumption.
         intros ;apply P_id_U21_monotonic;assumption.
         intros ;apply P_id_U31_monotonic;assumption.
         intros ;apply P_id_U42_monotonic;assumption.
         intros ;apply P_id_U22_monotonic;assumption.
         intros ;apply P_id_DIN_monotonic;assumption.
         intros ;apply P_id_U32_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_scc_4_large_scc_2_large_scc_2_strict_in_lt :
        Relation_Definitions.inclusion _ 
         DP_R_xml_0_scc_4_large_scc_2_large_scc_2_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_scc_4_large_scc_2_large_scc_2_large_in_le :
        Relation_Definitions.inclusion _ 
         DP_R_xml_0_scc_4_large_scc_2_large_scc_2_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_scc_4_large_scc_2_large_scc_2_large  := 
         WF_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large.wf.
       
       
       Lemma wf :
        well_founded WF_DP_R_xml_0_scc_4_large_scc_2_large.DP_R_xml_0_scc_4_large_scc_2_large_scc_2
         .
       Proof.
         intros x.
         apply (well_founded_ind wf_lt).
         clear x.
         intros x.
         pattern x.
         apply (@Acc_ind _ DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large).
         clear x.
         intros x _ IHx IHx'.
         constructor.
         intros y H.
         
         destruct H;
          (apply IHx';
            apply DP_R_xml_0_scc_4_large_scc_2_large_scc_2_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_scc_4_large_scc_2_large_scc_2_large_in_le;
             econstructor eassumption])).
         apply wf_DP_R_xml_0_scc_4_large_scc_2_large_scc_2_large.
       Qed.
      End WF_DP_R_xml_0_scc_4_large_scc_2_large_scc_2.
      
      Definition wf_DP_R_xml_0_scc_4_large_scc_2_large_scc_2  := 
        WF_DP_R_xml_0_scc_4_large_scc_2_large_scc_2.wf.
      
      
      Lemma acc_DP_R_xml_0_scc_4_large_scc_2_large_scc_2 :
       forall x y, 
        (DP_R_xml_0_scc_4_large_scc_2_large_scc_2 x y) ->
         Acc WF_DP_R_xml_0_scc_4_large_scc_2.DP_R_xml_0_scc_4_large_scc_2_large 
          x.
      Proof.
        intros x.
        pattern x.
        apply (@Acc_ind _ DP_R_xml_0_scc_4_large_scc_2_large_scc_2).
        intros x' _ Hrec y h.
        
        inversion h;clear h;subst;
         constructor;intros _y _h;inversion _h;clear _h;subst;
          (eapply Hrec;econstructor eassumption)||
          ((eapply acc_DP_R_xml_0_scc_4_large_scc_2_large_non_scc_1;
             econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
           ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||
            (eapply Hrec;
              econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))).
        apply wf_DP_R_xml_0_scc_4_large_scc_2_large_scc_2.
      Qed.
      
      
      Lemma wf :
       well_founded WF_DP_R_xml_0_scc_4_large_scc_2.DP_R_xml_0_scc_4_large_scc_2_large
        .
      Proof.
        constructor;intros _y _h;inversion _h;clear _h;subst;
         (eapply acc_DP_R_xml_0_scc_4_large_scc_2_large_non_scc_1;
           econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
         ((eapply acc_DP_R_xml_0_scc_4_large_scc_2_large_non_scc_0;
            econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
          ((eapply acc_DP_R_xml_0_scc_4_large_scc_2_large_scc_2;
             econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
           ((eapply acc_DP_R_xml_0_scc_4_large_scc_2_large_scc_1;
              econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
            ((eapply acc_DP_R_xml_0_scc_4_large_scc_2_large_scc_0;
               econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
             ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||(fail)))))).
      Qed.
     End WF_DP_R_xml_0_scc_4_large_scc_2_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_din (x7:Z) := 0.
     
     Definition P_id_u32 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 3* x7 + 3* x10.
     
     Definition P_id_dout (x7:Z) := 2 + 1* x7.
     
     Definition P_id_plus (x7:Z) (x8:Z) := 3.
     
     Definition P_id_u42 (x7:Z) (x8:Z) (x9:Z) := 1 + 2* x7 + 1* x9.
     
     Definition P_id_times (x7:Z) (x8:Z) := 0.
     
     Definition P_id_der (x7:Z) := 0.
     
     Definition P_id_u41 (x7:Z) (x8:Z) := 2* x7.
     
     Definition P_id_u22 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 2 + 3* x7 + 1* x10.
     
     Definition P_id_u21 (x7:Z) (x8:Z) (x9:Z) := 2* x7.
     
     Definition P_id_u31 (x7:Z) (x8:Z) (x9:Z) := 3* x7.
     
     Lemma P_id_din_monotonic :
      forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_din x8 <= P_id_din x7.
     Proof.
       intros x8 x7.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_u32_monotonic :
      forall x8 x12 x10 x14 x9 x13 x11 x7, 
       (0 <= x14)/\ (x14 <= x13) ->
        (0 <= x12)/\ (x12 <= x11) ->
         (0 <= x10)/\ (x10 <= x9) ->
          (0 <= x8)/\ (x8 <= x7) ->
           P_id_u32 x8 x10 x12 x14 <= P_id_u32 x7 x9 x11 x13.
     Proof.
       intros x14 x13 x12 x11 x10 x9 x8 x7.
       intros [H_1 H_0].
       intros [H_3 H_2].
       intros [H_5 H_4].
       intros [H_7 H_6].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_dout_monotonic :
      forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_dout x8 <= P_id_dout x7.
     Proof.
       intros x8 x7.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_plus_monotonic :
      forall x8 x10 x9 x7, 
       (0 <= x10)/\ (x10 <= x9) ->
        (0 <= x8)/\ (x8 <= x7) ->P_id_plus x8 x10 <= P_id_plus x7 x9.
     Proof.
       intros x10 x9 x8 x7.
       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_u42_monotonic :
      forall x8 x12 x10 x9 x11 x7, 
       (0 <= x12)/\ (x12 <= x11) ->
        (0 <= x10)/\ (x10 <= x9) ->
         (0 <= x8)/\ (x8 <= x7) ->P_id_u42 x8 x10 x12 <= P_id_u42 x7 x9 x11.
     Proof.
       intros x12 x11 x10 x9 x8 x7.
       intros [H_1 H_0].
       intros [H_3 H_2].
       intros [H_5 H_4].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_times_monotonic :
      forall x8 x10 x9 x7, 
       (0 <= x10)/\ (x10 <= x9) ->
        (0 <= x8)/\ (x8 <= x7) ->P_id_times x8 x10 <= P_id_times x7 x9.
     Proof.
       intros x10 x9 x8 x7.
       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_der_monotonic :
      forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_der x8 <= P_id_der x7.
     Proof.
       intros x8 x7.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_u41_monotonic :
      forall x8 x10 x9 x7, 
       (0 <= x10)/\ (x10 <= x9) ->
        (0 <= x8)/\ (x8 <= x7) ->P_id_u41 x8 x10 <= P_id_u41 x7 x9.
     Proof.
       intros x10 x9 x8 x7.
       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_u22_monotonic :
      forall x8 x12 x10 x14 x9 x13 x11 x7, 
       (0 <= x14)/\ (x14 <= x13) ->
        (0 <= x12)/\ (x12 <= x11) ->
         (0 <= x10)/\ (x10 <= x9) ->
          (0 <= x8)/\ (x8 <= x7) ->
           P_id_u22 x8 x10 x12 x14 <= P_id_u22 x7 x9 x11 x13.
     Proof.
       intros x14 x13 x12 x11 x10 x9 x8 x7.
       intros [H_1 H_0].
       intros [H_3 H_2].
       intros [H_5 H_4].
       intros [H_7 H_6].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_u21_monotonic :
      forall x8 x12 x10 x9 x11 x7, 
       (0 <= x12)/\ (x12 <= x11) ->
        (0 <= x10)/\ (x10 <= x9) ->
         (0 <= x8)/\ (x8 <= x7) ->P_id_u21 x8 x10 x12 <= P_id_u21 x7 x9 x11.
     Proof.
       intros x12 x11 x10 x9 x8 x7.
       intros [H_1 H_0].
       intros [H_3 H_2].
       intros [H_5 H_4].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_u31_monotonic :
      forall x8 x12 x10 x9 x11 x7, 
       (0 <= x12)/\ (x12 <= x11) ->
        (0 <= x10)/\ (x10 <= x9) ->
         (0 <= x8)/\ (x8 <= x7) ->P_id_u31 x8 x10 x12 <= P_id_u31 x7 x9 x11.
     Proof.
       intros x12 x11 x10 x9 x8 x7.
       intros [H_1 H_0].
       intros [H_3 H_2].
       intros [H_5 H_4].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_din_bounded : forall x7, (0 <= x7) ->0 <= P_id_din x7.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_u32_bounded :
      forall x8 x10 x9 x7, 
       (0 <= x7) ->
        (0 <= x8) ->(0 <= x9) ->(0 <= x10) ->0 <= P_id_u32 x10 x9 x8 x7.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_dout_bounded : forall x7, (0 <= x7) ->0 <= P_id_dout x7.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_plus_bounded :
      forall x8 x7, (0 <= x7) ->(0 <= x8) ->0 <= P_id_plus x8 x7.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_u42_bounded :
      forall x8 x9 x7, 
       (0 <= x7) ->(0 <= x8) ->(0 <= x9) ->0 <= P_id_u42 x9 x8 x7.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_times_bounded :
      forall x8 x7, (0 <= x7) ->(0 <= x8) ->0 <= P_id_times x8 x7.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_der_bounded : forall x7, (0 <= x7) ->0 <= P_id_der x7.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_u41_bounded :
      forall x8 x7, (0 <= x7) ->(0 <= x8) ->0 <= P_id_u41 x8 x7.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_u22_bounded :
      forall x8 x10 x9 x7, 
       (0 <= x7) ->
        (0 <= x8) ->(0 <= x9) ->(0 <= x10) ->0 <= P_id_u22 x10 x9 x8 x7.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_u21_bounded :
      forall x8 x9 x7, 
       (0 <= x7) ->(0 <= x8) ->(0 <= x9) ->0 <= P_id_u21 x9 x8 x7.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_u31_bounded :
      forall x8 x9 x7, 
       (0 <= x7) ->(0 <= x8) ->(0 <= x9) ->0 <= P_id_u31 x9 x8 x7.
     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_din P_id_u32 P_id_dout P_id_plus P_id_u42 
        P_id_times P_id_der P_id_u41 P_id_u22 P_id_u21 P_id_u31.
     
     Lemma measure_equation :
      forall t, 
       measure t = match t with
                     | (algebra.Alg.Term algebra.F.id_din (x7::nil)) =>
                      P_id_din (measure x7)
                     | (algebra.Alg.Term algebra.F.id_u32 (x10::x9::x8::
                        x7::nil)) =>
                      P_id_u32 (measure x10) (measure x9) (measure x8) 
                       (measure x7)
                     | (algebra.Alg.Term algebra.F.id_dout (x7::nil)) =>
                      P_id_dout (measure x7)
                     | (algebra.Alg.Term algebra.F.id_plus (x8::x7::nil)) =>
                      P_id_plus (measure x8) (measure x7)
                     | (algebra.Alg.Term algebra.F.id_u42 (x9::x8::x7::nil)) =>
                      P_id_u42 (measure x9) (measure x8) (measure x7)
                     | (algebra.Alg.Term algebra.F.id_times (x8::x7::nil)) =>
                      P_id_times (measure x8) (measure x7)
                     | (algebra.Alg.Term algebra.F.id_der (x7::nil)) =>
                      P_id_der (measure x7)
                     | (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil)) =>
                      P_id_u41 (measure x8) (measure x7)
                     | (algebra.Alg.Term algebra.F.id_u22 (x10::x9::x8::
                        x7::nil)) =>
                      P_id_u22 (measure x10) (measure x9) (measure x8) 
                       (measure x7)
                     | (algebra.Alg.Term algebra.F.id_u21 (x9::x8::x7::nil)) =>
                      P_id_u21 (measure x9) (measure x8) (measure x7)
                     | (algebra.Alg.Term algebra.F.id_u31 (x9::x8::x7::nil)) =>
                      P_id_u31 (measure x9) (measure x8) (measure x7)
                     | _ => 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_din_monotonic;assumption.
       intros ;apply P_id_u32_monotonic;assumption.
       intros ;apply P_id_dout_monotonic;assumption.
       intros ;apply P_id_plus_monotonic;assumption.
       intros ;apply P_id_u42_monotonic;assumption.
       intros ;apply P_id_times_monotonic;assumption.
       intros ;apply P_id_der_monotonic;assumption.
       intros ;apply P_id_u41_monotonic;assumption.
       intros ;apply P_id_u22_monotonic;assumption.
       intros ;apply P_id_u21_monotonic;assumption.
       intros ;apply P_id_u31_monotonic;assumption.
       intros ;apply P_id_din_bounded;assumption.
       intros ;apply P_id_u32_bounded;assumption.
       intros ;apply P_id_dout_bounded;assumption.
       intros ;apply P_id_plus_bounded;assumption.
       intros ;apply P_id_u42_bounded;assumption.
       intros ;apply P_id_times_bounded;assumption.
       intros ;apply P_id_der_bounded;assumption.
       intros ;apply P_id_u41_bounded;assumption.
       intros ;apply P_id_u22_bounded;assumption.
       intros ;apply P_id_u21_bounded;assumption.
       intros ;apply P_id_u31_bounded;assumption.
       apply rules_monotonic.
     Qed.
     
     Definition P_id_U41 (x7:Z) (x8:Z) := 0.
     
     Definition P_id_U21 (x7:Z) (x8:Z) (x9:Z) := 0.
     
     Definition P_id_U31 (x7:Z) (x8:Z) (x9:Z) := 1* x7.
     
     Definition P_id_U42 (x7:Z) (x8:Z) (x9:Z) := 0.
     
     Definition P_id_U22 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 0.
     
     Definition P_id_DIN (x7:Z) := 0.
     
     Definition P_id_U32 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 0.
     
     Lemma P_id_U41_monotonic :
      forall x8 x10 x9 x7, 
       (0 <= x10)/\ (x10 <= x9) ->
        (0 <= x8)/\ (x8 <= x7) ->P_id_U41 x8 x10 <= P_id_U41 x7 x9.
     Proof.
       intros x10 x9 x8 x7.
       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_U21_monotonic :
      forall x8 x12 x10 x9 x11 x7, 
       (0 <= x12)/\ (x12 <= x11) ->
        (0 <= x10)/\ (x10 <= x9) ->
         (0 <= x8)/\ (x8 <= x7) ->P_id_U21 x8 x10 x12 <= P_id_U21 x7 x9 x11.
     Proof.
       intros x12 x11 x10 x9 x8 x7.
       intros [H_1 H_0].
       intros [H_3 H_2].
       intros [H_5 H_4].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U31_monotonic :
      forall x8 x12 x10 x9 x11 x7, 
       (0 <= x12)/\ (x12 <= x11) ->
        (0 <= x10)/\ (x10 <= x9) ->
         (0 <= x8)/\ (x8 <= x7) ->P_id_U31 x8 x10 x12 <= P_id_U31 x7 x9 x11.
     Proof.
       intros x12 x11 x10 x9 x8 x7.
       intros [H_1 H_0].
       intros [H_3 H_2].
       intros [H_5 H_4].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U42_monotonic :
      forall x8 x12 x10 x9 x11 x7, 
       (0 <= x12)/\ (x12 <= x11) ->
        (0 <= x10)/\ (x10 <= x9) ->
         (0 <= x8)/\ (x8 <= x7) ->P_id_U42 x8 x10 x12 <= P_id_U42 x7 x9 x11.
     Proof.
       intros x12 x11 x10 x9 x8 x7.
       intros [H_1 H_0].
       intros [H_3 H_2].
       intros [H_5 H_4].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U22_monotonic :
      forall x8 x12 x10 x14 x9 x13 x11 x7, 
       (0 <= x14)/\ (x14 <= x13) ->
        (0 <= x12)/\ (x12 <= x11) ->
         (0 <= x10)/\ (x10 <= x9) ->
          (0 <= x8)/\ (x8 <= x7) ->
           P_id_U22 x8 x10 x12 x14 <= P_id_U22 x7 x9 x11 x13.
     Proof.
       intros x14 x13 x12 x11 x10 x9 x8 x7.
       intros [H_1 H_0].
       intros [H_3 H_2].
       intros [H_5 H_4].
       intros [H_7 H_6].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_DIN_monotonic :
      forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_DIN x8 <= P_id_DIN x7.
     Proof.
       intros x8 x7.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U32_monotonic :
      forall x8 x12 x10 x14 x9 x13 x11 x7, 
       (0 <= x14)/\ (x14 <= x13) ->
        (0 <= x12)/\ (x12 <= x11) ->
         (0 <= x10)/\ (x10 <= x9) ->
          (0 <= x8)/\ (x8 <= x7) ->
           P_id_U32 x8 x10 x12 x14 <= P_id_U32 x7 x9 x11 x13.
     Proof.
       intros x14 x13 x12 x11 x10 x9 x8 x7.
       intros [H_1 H_0].
       intros [H_3 H_2].
       intros [H_5 H_4].
       intros [H_7 H_6].
       
       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_din P_id_u32 P_id_dout P_id_plus 
        P_id_u42 P_id_times P_id_der P_id_u41 P_id_u22 P_id_u21 P_id_u31 
        P_id_U41 P_id_U21 P_id_U31 P_id_U42 P_id_U22 P_id_DIN P_id_U32.
     
     Lemma marked_measure_equation :
      forall t, 
       marked_measure t = match t with
                            | (algebra.Alg.Term algebra.F.id_u41 (x8::
                               x7::nil)) =>
                             P_id_U41 (measure x8) (measure x7)
                            | (algebra.Alg.Term algebra.F.id_u21 (x9::x8::
                               x7::nil)) =>
                             P_id_U21 (measure x9) (measure x8) (measure x7)
                            | (algebra.Alg.Term algebra.F.id_u31 (x9::x8::
                               x7::nil)) =>
                             P_id_U31 (measure x9) (measure x8) (measure x7)
                            | (algebra.Alg.Term algebra.F.id_u42 (x9::x8::
                               x7::nil)) =>
                             P_id_U42 (measure x9) (measure x8) (measure x7)
                            | (algebra.Alg.Term algebra.F.id_u22 (x10::x9::
                               x8::x7::nil)) =>
                             P_id_U22 (measure x10) (measure x9) 
                              (measure x8) (measure x7)
                            | (algebra.Alg.Term algebra.F.id_din (x7::nil)) =>
                             P_id_DIN (measure x7)
                            | (algebra.Alg.Term algebra.F.id_u32 (x10::x9::
                               x8::x7::nil)) =>
                             P_id_U32 (measure x10) (measure x9) 
                              (measure x8) (measure x7)
                            | _ => 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_din_monotonic;assumption.
       intros ;apply P_id_u32_monotonic;assumption.
       intros ;apply P_id_dout_monotonic;assumption.
       intros ;apply P_id_plus_monotonic;assumption.
       intros ;apply P_id_u42_monotonic;assumption.
       intros ;apply P_id_times_monotonic;assumption.
       intros ;apply P_id_der_monotonic;assumption.
       intros ;apply P_id_u41_monotonic;assumption.
       intros ;apply P_id_u22_monotonic;assumption.
       intros ;apply P_id_u21_monotonic;assumption.
       intros ;apply P_id_u31_monotonic;assumption.
       intros ;apply P_id_din_bounded;assumption.
       intros ;apply P_id_u32_bounded;assumption.
       intros ;apply P_id_dout_bounded;assumption.
       intros ;apply P_id_plus_bounded;assumption.
       intros ;apply P_id_u42_bounded;assumption.
       intros ;apply P_id_times_bounded;assumption.
       intros ;apply P_id_der_bounded;assumption.
       intros ;apply P_id_u41_bounded;assumption.
       intros ;apply P_id_u22_bounded;assumption.
       intros ;apply P_id_u21_bounded;assumption.
       intros ;apply P_id_u31_bounded;assumption.
       apply rules_monotonic.
       intros ;apply P_id_U41_monotonic;assumption.
       intros ;apply P_id_U21_monotonic;assumption.
       intros ;apply P_id_U31_monotonic;assumption.
       intros ;apply P_id_U42_monotonic;assumption.
       intros ;apply P_id_U22_monotonic;assumption.
       intros ;apply P_id_DIN_monotonic;assumption.
       intros ;apply P_id_U32_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_scc_4_large_scc_2_strict_in_lt :
      Relation_Definitions.inclusion _ DP_R_xml_0_scc_4_large_scc_2_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_scc_4_large_scc_2_large_in_le :
      Relation_Definitions.inclusion _ DP_R_xml_0_scc_4_large_scc_2_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_scc_4_large_scc_2_large  := 
       WF_DP_R_xml_0_scc_4_large_scc_2_large.wf.
     
     
     Lemma wf :
      well_founded WF_DP_R_xml_0_scc_4_large.DP_R_xml_0_scc_4_large_scc_2.
     Proof.
       intros x.
       apply (well_founded_ind wf_lt).
       clear x.
       intros x.
       pattern x.
       apply (@Acc_ind _ DP_R_xml_0_scc_4_large_scc_2_large).
       clear x.
       intros x _ IHx IHx'.
       constructor.
       intros y H.
       
       destruct H;
        (apply IHx';apply DP_R_xml_0_scc_4_large_scc_2_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_scc_4_large_scc_2_large_in_le;
           econstructor eassumption])).
       apply wf_DP_R_xml_0_scc_4_large_scc_2_large.
     Qed.
    End WF_DP_R_xml_0_scc_4_large_scc_2.
    
    Definition wf_DP_R_xml_0_scc_4_large_scc_2  := 
      WF_DP_R_xml_0_scc_4_large_scc_2.wf.
    
    
    Lemma acc_DP_R_xml_0_scc_4_large_scc_2 :
     forall x y, 
      (DP_R_xml_0_scc_4_large_scc_2 x y) ->
       Acc WF_DP_R_xml_0_scc_4.DP_R_xml_0_scc_4_large x.
    Proof.
      intros x.
      pattern x.
      apply (@Acc_ind _ DP_R_xml_0_scc_4_large_scc_2).
      intros x' _ Hrec y h.
      
      inversion h;clear h;subst;
       constructor;intros _y _h;inversion _h;clear _h;subst;
        (eapply Hrec;econstructor eassumption)||
        ((eapply acc_DP_R_xml_0_scc_4_large_non_scc_1;
           econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
         ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||
          (eapply Hrec;
            econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))).
      apply wf_DP_R_xml_0_scc_4_large_scc_2.
    Qed.
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_4.DP_R_xml_0_scc_4_large.
    Proof.
      constructor;intros _y _h;inversion _h;clear _h;subst;
       (eapply acc_DP_R_xml_0_scc_4_large_non_scc_1;
         econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
       ((eapply acc_DP_R_xml_0_scc_4_large_non_scc_0;
          econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
        ((eapply acc_DP_R_xml_0_scc_4_large_scc_2;
           econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
         ((eapply acc_DP_R_xml_0_scc_4_large_scc_1;
            econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
          ((eapply acc_DP_R_xml_0_scc_4_large_scc_0;
             econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
           ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||(fail)))))).
    Qed.
   End WF_DP_R_xml_0_scc_4_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_din (x7:Z) := 0.
   
   Definition P_id_u32 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 1* x7.
   
   Definition P_id_dout (x7:Z) := 2.
   
   Definition P_id_plus (x7:Z) (x8:Z) := 0.
   
   Definition P_id_u42 (x7:Z) (x8:Z) (x9:Z) := 3 + 3* x7.
   
   Definition P_id_times (x7:Z) (x8:Z) := 0.
   
   Definition P_id_der (x7:Z) := 0.
   
   Definition P_id_u41 (x7:Z) (x8:Z) := 2* x7.
   
   Definition P_id_u22 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 3 + 3* x7.
   
   Definition P_id_u21 (x7:Z) (x8:Z) (x9:Z) := 2* x7.
   
   Definition P_id_u31 (x7:Z) (x8:Z) (x9:Z) := 2* x7.
   
   Lemma P_id_din_monotonic :
    forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_din x8 <= P_id_din x7.
   Proof.
     intros x8 x7.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u32_monotonic :
    forall x8 x12 x10 x14 x9 x13 x11 x7, 
     (0 <= x14)/\ (x14 <= x13) ->
      (0 <= x12)/\ (x12 <= x11) ->
       (0 <= x10)/\ (x10 <= x9) ->
        (0 <= x8)/\ (x8 <= x7) ->
         P_id_u32 x8 x10 x12 x14 <= P_id_u32 x7 x9 x11 x13.
   Proof.
     intros x14 x13 x12 x11 x10 x9 x8 x7.
     intros [H_1 H_0].
     intros [H_3 H_2].
     intros [H_5 H_4].
     intros [H_7 H_6].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_dout_monotonic :
    forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_dout x8 <= P_id_dout x7.
   Proof.
     intros x8 x7.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_plus_monotonic :
    forall x8 x10 x9 x7, 
     (0 <= x10)/\ (x10 <= x9) ->
      (0 <= x8)/\ (x8 <= x7) ->P_id_plus x8 x10 <= P_id_plus x7 x9.
   Proof.
     intros x10 x9 x8 x7.
     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_u42_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (0 <= x12)/\ (x12 <= x11) ->
      (0 <= x10)/\ (x10 <= x9) ->
       (0 <= x8)/\ (x8 <= x7) ->P_id_u42 x8 x10 x12 <= P_id_u42 x7 x9 x11.
   Proof.
     intros x12 x11 x10 x9 x8 x7.
     intros [H_1 H_0].
     intros [H_3 H_2].
     intros [H_5 H_4].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_times_monotonic :
    forall x8 x10 x9 x7, 
     (0 <= x10)/\ (x10 <= x9) ->
      (0 <= x8)/\ (x8 <= x7) ->P_id_times x8 x10 <= P_id_times x7 x9.
   Proof.
     intros x10 x9 x8 x7.
     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_der_monotonic :
    forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_der x8 <= P_id_der x7.
   Proof.
     intros x8 x7.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u41_monotonic :
    forall x8 x10 x9 x7, 
     (0 <= x10)/\ (x10 <= x9) ->
      (0 <= x8)/\ (x8 <= x7) ->P_id_u41 x8 x10 <= P_id_u41 x7 x9.
   Proof.
     intros x10 x9 x8 x7.
     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_u22_monotonic :
    forall x8 x12 x10 x14 x9 x13 x11 x7, 
     (0 <= x14)/\ (x14 <= x13) ->
      (0 <= x12)/\ (x12 <= x11) ->
       (0 <= x10)/\ (x10 <= x9) ->
        (0 <= x8)/\ (x8 <= x7) ->
         P_id_u22 x8 x10 x12 x14 <= P_id_u22 x7 x9 x11 x13.
   Proof.
     intros x14 x13 x12 x11 x10 x9 x8 x7.
     intros [H_1 H_0].
     intros [H_3 H_2].
     intros [H_5 H_4].
     intros [H_7 H_6].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u21_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (0 <= x12)/\ (x12 <= x11) ->
      (0 <= x10)/\ (x10 <= x9) ->
       (0 <= x8)/\ (x8 <= x7) ->P_id_u21 x8 x10 x12 <= P_id_u21 x7 x9 x11.
   Proof.
     intros x12 x11 x10 x9 x8 x7.
     intros [H_1 H_0].
     intros [H_3 H_2].
     intros [H_5 H_4].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u31_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (0 <= x12)/\ (x12 <= x11) ->
      (0 <= x10)/\ (x10 <= x9) ->
       (0 <= x8)/\ (x8 <= x7) ->P_id_u31 x8 x10 x12 <= P_id_u31 x7 x9 x11.
   Proof.
     intros x12 x11 x10 x9 x8 x7.
     intros [H_1 H_0].
     intros [H_3 H_2].
     intros [H_5 H_4].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_din_bounded : forall x7, (0 <= x7) ->0 <= P_id_din x7.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u32_bounded :
    forall x8 x10 x9 x7, 
     (0 <= x7) ->
      (0 <= x8) ->(0 <= x9) ->(0 <= x10) ->0 <= P_id_u32 x10 x9 x8 x7.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_dout_bounded : forall x7, (0 <= x7) ->0 <= P_id_dout x7.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_plus_bounded :
    forall x8 x7, (0 <= x7) ->(0 <= x8) ->0 <= P_id_plus x8 x7.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u42_bounded :
    forall x8 x9 x7, 
     (0 <= x7) ->(0 <= x8) ->(0 <= x9) ->0 <= P_id_u42 x9 x8 x7.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_times_bounded :
    forall x8 x7, (0 <= x7) ->(0 <= x8) ->0 <= P_id_times x8 x7.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_der_bounded : forall x7, (0 <= x7) ->0 <= P_id_der x7.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u41_bounded :
    forall x8 x7, (0 <= x7) ->(0 <= x8) ->0 <= P_id_u41 x8 x7.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u22_bounded :
    forall x8 x10 x9 x7, 
     (0 <= x7) ->
      (0 <= x8) ->(0 <= x9) ->(0 <= x10) ->0 <= P_id_u22 x10 x9 x8 x7.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u21_bounded :
    forall x8 x9 x7, 
     (0 <= x7) ->(0 <= x8) ->(0 <= x9) ->0 <= P_id_u21 x9 x8 x7.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u31_bounded :
    forall x8 x9 x7, 
     (0 <= x7) ->(0 <= x8) ->(0 <= x9) ->0 <= P_id_u31 x9 x8 x7.
   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_din P_id_u32 P_id_dout P_id_plus P_id_u42 
      P_id_times P_id_der P_id_u41 P_id_u22 P_id_u21 P_id_u31.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_din (x7::nil)) =>
                    P_id_din (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u32 (x10::x9::x8::
                      x7::nil)) =>
                    P_id_u32 (measure x10) (measure x9) (measure x8) 
                     (measure x7)
                   | (algebra.Alg.Term algebra.F.id_dout (x7::nil)) =>
                    P_id_dout (measure x7)
                   | (algebra.Alg.Term algebra.F.id_plus (x8::x7::nil)) =>
                    P_id_plus (measure x8) (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u42 (x9::x8::x7::nil)) =>
                    P_id_u42 (measure x9) (measure x8) (measure x7)
                   | (algebra.Alg.Term algebra.F.id_times (x8::x7::nil)) =>
                    P_id_times (measure x8) (measure x7)
                   | (algebra.Alg.Term algebra.F.id_der (x7::nil)) =>
                    P_id_der (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil)) =>
                    P_id_u41 (measure x8) (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u22 (x10::x9::x8::
                      x7::nil)) =>
                    P_id_u22 (measure x10) (measure x9) (measure x8) 
                     (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u21 (x9::x8::x7::nil)) =>
                    P_id_u21 (measure x9) (measure x8) (measure x7)
                   | (algebra.Alg.Term algebra.F.id_u31 (x9::x8::x7::nil)) =>
                    P_id_u31 (measure x9) (measure x8) (measure x7)
                   | _ => 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_din_monotonic;assumption.
     intros ;apply P_id_u32_monotonic;assumption.
     intros ;apply P_id_dout_monotonic;assumption.
     intros ;apply P_id_plus_monotonic;assumption.
     intros ;apply P_id_u42_monotonic;assumption.
     intros ;apply P_id_times_monotonic;assumption.
     intros ;apply P_id_der_monotonic;assumption.
     intros ;apply P_id_u41_monotonic;assumption.
     intros ;apply P_id_u22_monotonic;assumption.
     intros ;apply P_id_u21_monotonic;assumption.
     intros ;apply P_id_u31_monotonic;assumption.
     intros ;apply P_id_din_bounded;assumption.
     intros ;apply P_id_u32_bounded;assumption.
     intros ;apply P_id_dout_bounded;assumption.
     intros ;apply P_id_plus_bounded;assumption.
     intros ;apply P_id_u42_bounded;assumption.
     intros ;apply P_id_times_bounded;assumption.
     intros ;apply P_id_der_bounded;assumption.
     intros ;apply P_id_u41_bounded;assumption.
     intros ;apply P_id_u22_bounded;assumption.
     intros ;apply P_id_u21_bounded;assumption.
     intros ;apply P_id_u31_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U41 (x7:Z) (x8:Z) := 0.
   
   Definition P_id_U21 (x7:Z) (x8:Z) (x9:Z) := 2* x7.
   
   Definition P_id_U31 (x7:Z) (x8:Z) (x9:Z) := 0.
   
   Definition P_id_U42 (x7:Z) (x8:Z) (x9:Z) := 0.
   
   Definition P_id_U22 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 0.
   
   Definition P_id_DIN (x7:Z) := 0.
   
   Definition P_id_U32 (x7:Z) (x8:Z) (x9:Z) (x10:Z) := 0.
   
   Lemma P_id_U41_monotonic :
    forall x8 x10 x9 x7, 
     (0 <= x10)/\ (x10 <= x9) ->
      (0 <= x8)/\ (x8 <= x7) ->P_id_U41 x8 x10 <= P_id_U41 x7 x9.
   Proof.
     intros x10 x9 x8 x7.
     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_U21_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (0 <= x12)/\ (x12 <= x11) ->
      (0 <= x10)/\ (x10 <= x9) ->
       (0 <= x8)/\ (x8 <= x7) ->P_id_U21 x8 x10 x12 <= P_id_U21 x7 x9 x11.
   Proof.
     intros x12 x11 x10 x9 x8 x7.
     intros [H_1 H_0].
     intros [H_3 H_2].
     intros [H_5 H_4].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (0 <= x12)/\ (x12 <= x11) ->
      (0 <= x10)/\ (x10 <= x9) ->
       (0 <= x8)/\ (x8 <= x7) ->P_id_U31 x8 x10 x12 <= P_id_U31 x7 x9 x11.
   Proof.
     intros x12 x11 x10 x9 x8 x7.
     intros [H_1 H_0].
     intros [H_3 H_2].
     intros [H_5 H_4].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x8 x12 x10 x9 x11 x7, 
     (0 <= x12)/\ (x12 <= x11) ->
      (0 <= x10)/\ (x10 <= x9) ->
       (0 <= x8)/\ (x8 <= x7) ->P_id_U42 x8 x10 x12 <= P_id_U42 x7 x9 x11.
   Proof.
     intros x12 x11 x10 x9 x8 x7.
     intros [H_1 H_0].
     intros [H_3 H_2].
     intros [H_5 H_4].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x8 x12 x10 x14 x9 x13 x11 x7, 
     (0 <= x14)/\ (x14 <= x13) ->
      (0 <= x12)/\ (x12 <= x11) ->
       (0 <= x10)/\ (x10 <= x9) ->
        (0 <= x8)/\ (x8 <= x7) ->
         P_id_U22 x8 x10 x12 x14 <= P_id_U22 x7 x9 x11 x13.
   Proof.
     intros x14 x13 x12 x11 x10 x9 x8 x7.
     intros [H_1 H_0].
     intros [H_3 H_2].
     intros [H_5 H_4].
     intros [H_7 H_6].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_DIN_monotonic :
    forall x8 x7, (0 <= x8)/\ (x8 <= x7) ->P_id_DIN x8 <= P_id_DIN x7.
   Proof.
     intros x8 x7.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U32_monotonic :
    forall x8 x12 x10 x14 x9 x13 x11 x7, 
     (0 <= x14)/\ (x14 <= x13) ->
      (0 <= x12)/\ (x12 <= x11) ->
       (0 <= x10)/\ (x10 <= x9) ->
        (0 <= x8)/\ (x8 <= x7) ->
         P_id_U32 x8 x10 x12 x14 <= P_id_U32 x7 x9 x11 x13.
   Proof.
     intros x14 x13 x12 x11 x10 x9 x8 x7.
     intros [H_1 H_0].
     intros [H_3 H_2].
     intros [H_5 H_4].
     intros [H_7 H_6].
     
     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_din P_id_u32 P_id_dout P_id_plus P_id_u42 
      P_id_times P_id_der P_id_u41 P_id_u22 P_id_u21 P_id_u31 P_id_U41 
      P_id_U21 P_id_U31 P_id_U42 P_id_U22 P_id_DIN P_id_U32.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_u41 (x8::x7::nil)) =>
                           P_id_U41 (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u21 (x9::x8::
                             x7::nil)) =>
                           P_id_U21 (measure x9) (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u31 (x9::x8::
                             x7::nil)) =>
                           P_id_U31 (measure x9) (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u42 (x9::x8::
                             x7::nil)) =>
                           P_id_U42 (measure x9) (measure x8) (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u22 (x10::x9::x8::
                             x7::nil)) =>
                           P_id_U22 (measure x10) (measure x9) (measure x8) 
                            (measure x7)
                          | (algebra.Alg.Term algebra.F.id_din (x7::nil)) =>
                           P_id_DIN (measure x7)
                          | (algebra.Alg.Term algebra.F.id_u32 (x10::x9::x8::
                             x7::nil)) =>
                           P_id_U32 (measure x10) (measure x9) (measure x8) 
                            (measure x7)
                          | _ => 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_din_monotonic;assumption.
     intros ;apply P_id_u32_monotonic;assumption.
     intros ;apply P_id_dout_monotonic;assumption.
     intros ;apply P_id_plus_monotonic;assumption.
     intros ;apply P_id_u42_monotonic;assumption.
     intros ;apply P_id_times_monotonic;assumption.
     intros ;apply P_id_der_monotonic;assumption.
     intros ;apply P_id_u41_monotonic;assumption.
     intros ;apply P_id_u22_monotonic;assumption.
     intros ;apply P_id_u21_monotonic;assumption.
     intros ;apply P_id_u31_monotonic;assumption.
     intros ;apply P_id_din_bounded;assumption.
     intros ;apply P_id_u32_bounded;assumption.
     intros ;apply P_id_dout_bounded;assumption.
     intros ;apply P_id_plus_bounded;assumption.
     intros ;apply P_id_u42_bounded;assumption.
     intros ;apply P_id_times_bounded;assumption.
     intros ;apply P_id_der_bounded;assumption.
     intros ;apply P_id_u41_bounded;assumption.
     intros ;apply P_id_u22_bounded;assumption.
     intros ;apply P_id_u21_bounded;assumption.
     intros ;apply P_id_u31_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_DIN_monotonic;assumption.
     intros ;apply P_id_U32_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_scc_4_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_4_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_scc_4_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_4_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_scc_4_large  := WF_DP_R_xml_0_scc_4_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_4.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_4_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_4_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_scc_4_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_4_large.
   Qed.
  End WF_DP_R_xml_0_scc_4.
  
  Definition wf_DP_R_xml_0_scc_4  := WF_DP_R_xml_0_scc_4.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_4 :
   forall x y, (DP_R_xml_0_scc_4 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x.
  Proof.
    intros x.
    pattern x.
    apply (@Acc_ind _ DP_R_xml_0_scc_4).
    intros x' _ Hrec y h.
    
    inversion h;clear h;subst;
     constructor;intros _y _h;inversion _h;clear _h;subst;
      (eapply Hrec;econstructor eassumption)||
      ((eapply acc_DP_R_xml_0_non_scc_3;
         econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
       ((eapply acc_DP_R_xml_0_non_scc_2;
          econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
        ((eapply acc_DP_R_xml_0_non_scc_1;
           econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
         ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||
          (eapply Hrec;
            econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))))).
    apply wf_DP_R_xml_0_scc_4.
  Qed.
  
  
  Lemma wf : well_founded WF_R_xml_0_deep_rew.DP_R_xml_0.
  Proof.
    constructor;intros _y _h;inversion _h;clear _h;subst;
     (eapply acc_DP_R_xml_0_non_scc_3;
       econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
     ((eapply acc_DP_R_xml_0_non_scc_2;
        econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
      ((eapply acc_DP_R_xml_0_non_scc_1;
         econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
       ((eapply acc_DP_R_xml_0_non_scc_0;
          econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
        ((eapply acc_DP_R_xml_0_scc_4;
           econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
         ((eapply acc_DP_R_xml_0_scc_3;
            econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
          ((eapply acc_DP_R_xml_0_scc_2;
             econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
           ((eapply acc_DP_R_xml_0_scc_1;
              econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
            ((eapply acc_DP_R_xml_0_scc_0;
               econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
             ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||(fail)))))))))).
  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___Rubio___wst99.trs/a3pat.v" ***
*** End: ***
 *)