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_active *)
    | id_active : symb
     (* id_add *)
    | id_add : symb
     (* id_recip *)
    | id_recip : symb
     (* id_mark *)
    | id_mark : symb
     (* id_first *)
    | id_first : symb
     (* id_s *)
    | id_s : symb
     (* id_terms *)
    | id_terms : symb
     (* id_dbl *)
    | id_dbl : symb
     (* id_sqr *)
    | id_sqr : symb
     (* id_cons *)
    | id_cons : symb
     (* id_nil *)
    | id_nil : symb
     (* id_0 *)
    | id_0 : symb
  .
  
  
  Definition symb_eq_bool (f1 f2:symb) : bool := 
    match f1,f2 with
      | id_active,id_active => true
      | id_add,id_add => true
      | id_recip,id_recip => true
      | id_mark,id_mark => true
      | id_first,id_first => true
      | id_s,id_s => true
      | id_terms,id_terms => true
      | id_dbl,id_dbl => true
      | id_sqr,id_sqr => true
      | id_cons,id_cons => true
      | id_nil,id_nil => true
      | id_0,id_0 => 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_active,id_active => refl_equal _
             | id_add,id_add => refl_equal _
             | id_recip,id_recip => refl_equal _
             | id_mark,id_mark => refl_equal _
             | id_first,id_first => refl_equal _
             | id_s,id_s => refl_equal _
             | id_terms,id_terms => refl_equal _
             | id_dbl,id_dbl => refl_equal _
             | id_sqr,id_sqr => refl_equal _
             | id_cons,id_cons => refl_equal _
             | id_nil,id_nil => refl_equal _
             | id_0,id_0 => refl_equal _
             | _,_ => _
             end;intros abs;discriminate.
  Defined.
  
  
  Definition arity (f:symb) := 
    match f with
      | id_active => term.Free 1
      | id_add => term.Free 2
      | id_recip => term.Free 1
      | id_mark => term.Free 1
      | id_first => term.Free 2
      | id_s => term.Free 1
      | id_terms => term.Free 1
      | id_dbl => term.Free 1
      | id_sqr => term.Free 1
      | id_cons => term.Free 2
      | id_nil => term.Free 0
      | id_0 => term.Free 0
      end.
  
  
  Definition symb_order (f1 f2:symb) : bool := 
    match f1,f2 with
      | id_active,id_active => true
      | id_active,id_add => false
      | id_active,id_recip => false
      | id_active,id_mark => false
      | id_active,id_first => false
      | id_active,id_s => false
      | id_active,id_terms => false
      | id_active,id_dbl => false
      | id_active,id_sqr => false
      | id_active,id_cons => false
      | id_active,id_nil => false
      | id_active,id_0 => false
      | id_add,id_active => true
      | id_add,id_add => true
      | id_add,id_recip => false
      | id_add,id_mark => false
      | id_add,id_first => false
      | id_add,id_s => false
      | id_add,id_terms => false
      | id_add,id_dbl => false
      | id_add,id_sqr => false
      | id_add,id_cons => false
      | id_add,id_nil => false
      | id_add,id_0 => false
      | id_recip,id_active => true
      | id_recip,id_add => true
      | id_recip,id_recip => true
      | id_recip,id_mark => false
      | id_recip,id_first => false
      | id_recip,id_s => false
      | id_recip,id_terms => false
      | id_recip,id_dbl => false
      | id_recip,id_sqr => false
      | id_recip,id_cons => false
      | id_recip,id_nil => false
      | id_recip,id_0 => false
      | id_mark,id_active => true
      | id_mark,id_add => true
      | id_mark,id_recip => true
      | id_mark,id_mark => true
      | id_mark,id_first => false
      | id_mark,id_s => false
      | id_mark,id_terms => false
      | id_mark,id_dbl => false
      | id_mark,id_sqr => false
      | id_mark,id_cons => false
      | id_mark,id_nil => false
      | id_mark,id_0 => false
      | id_first,id_active => true
      | id_first,id_add => true
      | id_first,id_recip => true
      | id_first,id_mark => true
      | id_first,id_first => true
      | id_first,id_s => false
      | id_first,id_terms => false
      | id_first,id_dbl => false
      | id_first,id_sqr => false
      | id_first,id_cons => false
      | id_first,id_nil => false
      | id_first,id_0 => false
      | id_s,id_active => true
      | id_s,id_add => true
      | id_s,id_recip => true
      | id_s,id_mark => true
      | id_s,id_first => true
      | id_s,id_s => true
      | id_s,id_terms => false
      | id_s,id_dbl => false
      | id_s,id_sqr => false
      | id_s,id_cons => false
      | id_s,id_nil => false
      | id_s,id_0 => false
      | id_terms,id_active => true
      | id_terms,id_add => true
      | id_terms,id_recip => true
      | id_terms,id_mark => true
      | id_terms,id_first => true
      | id_terms,id_s => true
      | id_terms,id_terms => true
      | id_terms,id_dbl => false
      | id_terms,id_sqr => false
      | id_terms,id_cons => false
      | id_terms,id_nil => false
      | id_terms,id_0 => false
      | id_dbl,id_active => true
      | id_dbl,id_add => true
      | id_dbl,id_recip => true
      | id_dbl,id_mark => true
      | id_dbl,id_first => true
      | id_dbl,id_s => true
      | id_dbl,id_terms => true
      | id_dbl,id_dbl => true
      | id_dbl,id_sqr => false
      | id_dbl,id_cons => false
      | id_dbl,id_nil => false
      | id_dbl,id_0 => false
      | id_sqr,id_active => true
      | id_sqr,id_add => true
      | id_sqr,id_recip => true
      | id_sqr,id_mark => true
      | id_sqr,id_first => true
      | id_sqr,id_s => true
      | id_sqr,id_terms => true
      | id_sqr,id_dbl => true
      | id_sqr,id_sqr => true
      | id_sqr,id_cons => false
      | id_sqr,id_nil => false
      | id_sqr,id_0 => false
      | id_cons,id_active => true
      | id_cons,id_add => true
      | id_cons,id_recip => true
      | id_cons,id_mark => true
      | id_cons,id_first => true
      | id_cons,id_s => true
      | id_cons,id_terms => true
      | id_cons,id_dbl => true
      | id_cons,id_sqr => true
      | id_cons,id_cons => true
      | id_cons,id_nil => false
      | id_cons,id_0 => false
      | id_nil,id_active => true
      | id_nil,id_add => true
      | id_nil,id_recip => true
      | id_nil,id_mark => true
      | id_nil,id_first => true
      | id_nil,id_s => true
      | id_nil,id_terms => true
      | id_nil,id_dbl => true
      | id_nil,id_sqr => true
      | id_nil,id_cons => true
      | id_nil,id_nil => true
      | id_nil,id_0 => false
      | id_0,id_active => true
      | id_0,id_add => true
      | id_0,id_recip => true
      | id_0,id_mark => true
      | id_0,id_first => true
      | id_0,id_s => true
      | id_0,id_terms => true
      | id_0,id_dbl => true
      | id_0,id_sqr => true
      | id_0,id_cons => true
      | id_0,id_nil => true
      | id_0,id_0 => 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 := 
    (* active(terms(N_)) -> mark(cons(recip(sqr(N_)),terms(s(N_)))) *)
   | R_xml_0_rule_0 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_recip 
                   ((algebra.Alg.Term algebra.F.id_sqr 
                   ((algebra.Alg.Var 1)::nil))::nil))::(algebra.Alg.Term 
                   algebra.F.id_terms ((algebra.Alg.Term algebra.F.id_s 
                   ((algebra.Alg.Var 1)::nil))::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_terms ((algebra.Alg.Var 1)::nil))::nil))
    (* active(sqr(0)) -> mark(0) *)
   | R_xml_0_rule_1 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_0 nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_0 nil)::nil))::nil))
   
    (* active(sqr(s(X_))) -> mark(s(add(sqr(X_),dbl(X_)))) *)
   | R_xml_0_rule_2 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_s ((algebra.Alg.Term algebra.F.id_add 
                   ((algebra.Alg.Term algebra.F.id_sqr 
                   ((algebra.Alg.Var 2)::nil))::(algebra.Alg.Term 
                   algebra.F.id_dbl 
                   ((algebra.Alg.Var 2)::nil))::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_s 
      ((algebra.Alg.Var 2)::nil))::nil))::nil))
    (* active(dbl(0)) -> mark(0) *)
   | R_xml_0_rule_3 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_0 nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_0 nil)::nil))::nil))
    (* active(dbl(s(X_))) -> mark(s(s(dbl(X_)))) *)
   | R_xml_0_rule_4 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_s ((algebra.Alg.Term algebra.F.id_s 
                   ((algebra.Alg.Term algebra.F.id_dbl 
                   ((algebra.Alg.Var 2)::nil))::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_s 
      ((algebra.Alg.Var 2)::nil))::nil))::nil))
    (* active(add(0,X_)) -> mark(X_) *)
   | R_xml_0_rule_5 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark 
                   ((algebra.Alg.Var 2)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_add ((algebra.Alg.Term algebra.F.id_0 nil)::
      (algebra.Alg.Var 2)::nil))::nil))
    (* active(add(s(X_),Y_)) -> mark(s(add(X_,Y_))) *)
   | R_xml_0_rule_6 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_s ((algebra.Alg.Term algebra.F.id_add 
                   ((algebra.Alg.Var 2)::
                   (algebra.Alg.Var 3)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_add ((algebra.Alg.Term algebra.F.id_s 
      ((algebra.Alg.Var 2)::nil))::(algebra.Alg.Var 3)::nil))::nil))
    (* active(first(0,X_)) -> mark(nil) *)
   | R_xml_0_rule_7 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_nil nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_first ((algebra.Alg.Term algebra.F.id_0 nil)::
      (algebra.Alg.Var 2)::nil))::nil))
   
    (* active(first(s(X_),cons(Y_,Z_))) -> mark(cons(Y_,first(X_,Z_))) *)
   | R_xml_0_rule_8 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_cons ((algebra.Alg.Var 3)::(algebra.Alg.Term 
                   algebra.F.id_first ((algebra.Alg.Var 2)::
                   (algebra.Alg.Var 4)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_first ((algebra.Alg.Term algebra.F.id_s 
      ((algebra.Alg.Var 2)::nil))::(algebra.Alg.Term algebra.F.id_cons 
      ((algebra.Alg.Var 3)::(algebra.Alg.Var 4)::nil))::nil))::nil))
    (* mark(terms(X_)) -> active(terms(mark(X_))) *)
   | R_xml_0_rule_9 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                   algebra.F.id_terms ((algebra.Alg.Term algebra.F.id_mark 
                   ((algebra.Alg.Var 2)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
      algebra.F.id_terms ((algebra.Alg.Var 2)::nil))::nil))
   
    (* mark(cons(X1_,X2_)) -> active(cons(mark(X1_),X2_)) *)
   | R_xml_0_rule_10 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                   algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_mark 
                   ((algebra.Alg.Var 5)::nil))::
                   (algebra.Alg.Var 6)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
      algebra.F.id_cons ((algebra.Alg.Var 5)::
      (algebra.Alg.Var 6)::nil))::nil))
    (* mark(recip(X_)) -> active(recip(mark(X_))) *)
   | R_xml_0_rule_11 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                   algebra.F.id_recip ((algebra.Alg.Term algebra.F.id_mark 
                   ((algebra.Alg.Var 2)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
      algebra.F.id_recip ((algebra.Alg.Var 2)::nil))::nil))
    (* mark(sqr(X_)) -> active(sqr(mark(X_))) *)
   | R_xml_0_rule_12 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                   algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_mark 
                   ((algebra.Alg.Var 2)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_sqr 
      ((algebra.Alg.Var 2)::nil))::nil))
    (* mark(s(X_)) -> active(s(X_)) *)
   | R_xml_0_rule_13 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                   algebra.F.id_s ((algebra.Alg.Var 2)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_s 
      ((algebra.Alg.Var 2)::nil))::nil))
    (* mark(0) -> active(0) *)
   | R_xml_0_rule_14 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                   algebra.F.id_0 nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_0 
      nil)::nil))
   
    (* mark(add(X1_,X2_)) -> active(add(mark(X1_),mark(X2_))) *)
   | R_xml_0_rule_15 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                   algebra.F.id_add ((algebra.Alg.Term algebra.F.id_mark 
                   ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Term 
                   algebra.F.id_mark 
                   ((algebra.Alg.Var 6)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_add 
      ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil))
    (* mark(dbl(X_)) -> active(dbl(mark(X_))) *)
   | R_xml_0_rule_16 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                   algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_mark 
                   ((algebra.Alg.Var 2)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_dbl 
      ((algebra.Alg.Var 2)::nil))::nil))
   
    (* mark(first(X1_,X2_)) -> active(first(mark(X1_),mark(X2_))) *)
   | R_xml_0_rule_17 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                   algebra.F.id_first ((algebra.Alg.Term algebra.F.id_mark 
                   ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Term 
                   algebra.F.id_mark 
                   ((algebra.Alg.Var 6)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
      algebra.F.id_first ((algebra.Alg.Var 5)::
      (algebra.Alg.Var 6)::nil))::nil))
    (* mark(nil) -> active(nil) *)
   | R_xml_0_rule_18 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                   algebra.F.id_nil nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_nil 
      nil)::nil))
    (* terms(mark(X_)) -> terms(X_) *)
   | R_xml_0_rule_19 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_terms 
                   ((algebra.Alg.Var 2)::nil)) 
     (algebra.Alg.Term algebra.F.id_terms ((algebra.Alg.Term 
      algebra.F.id_mark ((algebra.Alg.Var 2)::nil))::nil))
    (* terms(active(X_)) -> terms(X_) *)
   | R_xml_0_rule_20 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_terms 
                   ((algebra.Alg.Var 2)::nil)) 
     (algebra.Alg.Term algebra.F.id_terms ((algebra.Alg.Term 
      algebra.F.id_active ((algebra.Alg.Var 2)::nil))::nil))
    (* cons(mark(X1_),X2_) -> cons(X1_,X2_) *)
   | R_xml_0_rule_21 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::
                   (algebra.Alg.Var 6)::nil)) 
     (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term 
      algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::
      (algebra.Alg.Var 6)::nil))
    (* cons(X1_,mark(X2_)) -> cons(X1_,X2_) *)
   | R_xml_0_rule_22 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::
                   (algebra.Alg.Var 6)::nil)) 
     (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::
      (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 6)::nil))::nil))
    (* cons(active(X1_),X2_) -> cons(X1_,X2_) *)
   | R_xml_0_rule_23 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::
                   (algebra.Alg.Var 6)::nil)) 
     (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term 
      algebra.F.id_active ((algebra.Alg.Var 5)::nil))::
      (algebra.Alg.Var 6)::nil))
    (* cons(X1_,active(X2_)) -> cons(X1_,X2_) *)
   | R_xml_0_rule_24 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::
                   (algebra.Alg.Var 6)::nil)) 
     (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::
      (algebra.Alg.Term algebra.F.id_active 
      ((algebra.Alg.Var 6)::nil))::nil))
    (* recip(mark(X_)) -> recip(X_) *)
   | R_xml_0_rule_25 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_recip 
                   ((algebra.Alg.Var 2)::nil)) 
     (algebra.Alg.Term algebra.F.id_recip ((algebra.Alg.Term 
      algebra.F.id_mark ((algebra.Alg.Var 2)::nil))::nil))
    (* recip(active(X_)) -> recip(X_) *)
   | R_xml_0_rule_26 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_recip 
                   ((algebra.Alg.Var 2)::nil)) 
     (algebra.Alg.Term algebra.F.id_recip ((algebra.Alg.Term 
      algebra.F.id_active ((algebra.Alg.Var 2)::nil))::nil))
    (* sqr(mark(X_)) -> sqr(X_) *)
   | R_xml_0_rule_27 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_sqr 
                   ((algebra.Alg.Var 2)::nil)) 
     (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 2)::nil))::nil))
    (* sqr(active(X_)) -> sqr(X_) *)
   | R_xml_0_rule_28 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_sqr 
                   ((algebra.Alg.Var 2)::nil)) 
     (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term 
      algebra.F.id_active ((algebra.Alg.Var 2)::nil))::nil))
    (* s(mark(X_)) -> s(X_) *)
   | R_xml_0_rule_29 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_s 
                   ((algebra.Alg.Var 2)::nil)) 
     (algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 2)::nil))::nil))
    (* s(active(X_)) -> s(X_) *)
   | R_xml_0_rule_30 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_s 
                   ((algebra.Alg.Var 2)::nil)) 
     (algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_active 
      ((algebra.Alg.Var 2)::nil))::nil))
    (* add(mark(X1_),X2_) -> add(X1_,X2_) *)
   | R_xml_0_rule_31 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Var 5)::
                   (algebra.Alg.Var 6)::nil)) 
     (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Var 6)::nil))
    (* add(X1_,mark(X2_)) -> add(X1_,X2_) *)
   | R_xml_0_rule_32 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Var 5)::
                   (algebra.Alg.Var 6)::nil)) 
     (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Var 5)::
      (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 6)::nil))::nil))
    (* add(active(X1_),X2_) -> add(X1_,X2_) *)
   | R_xml_0_rule_33 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Var 5)::
                   (algebra.Alg.Var 6)::nil)) 
     (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term 
      algebra.F.id_active ((algebra.Alg.Var 5)::nil))::
      (algebra.Alg.Var 6)::nil))
    (* add(X1_,active(X2_)) -> add(X1_,X2_) *)
   | R_xml_0_rule_34 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Var 5)::
                   (algebra.Alg.Var 6)::nil)) 
     (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Var 5)::
      (algebra.Alg.Term algebra.F.id_active 
      ((algebra.Alg.Var 6)::nil))::nil))
    (* dbl(mark(X_)) -> dbl(X_) *)
   | R_xml_0_rule_35 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_dbl 
                   ((algebra.Alg.Var 2)::nil)) 
     (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 2)::nil))::nil))
    (* dbl(active(X_)) -> dbl(X_) *)
   | R_xml_0_rule_36 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_dbl 
                   ((algebra.Alg.Var 2)::nil)) 
     (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term 
      algebra.F.id_active ((algebra.Alg.Var 2)::nil))::nil))
    (* first(mark(X1_),X2_) -> first(X1_,X2_) *)
   | R_xml_0_rule_37 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Var 5)::
                   (algebra.Alg.Var 6)::nil)) 
     (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Term 
      algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::
      (algebra.Alg.Var 6)::nil))
    (* first(X1_,mark(X2_)) -> first(X1_,X2_) *)
   | R_xml_0_rule_38 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Var 5)::
                   (algebra.Alg.Var 6)::nil)) 
     (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Var 5)::
      (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 6)::nil))::nil))
    (* first(active(X1_),X2_) -> first(X1_,X2_) *)
   | R_xml_0_rule_39 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Var 5)::
                   (algebra.Alg.Var 6)::nil)) 
     (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Term 
      algebra.F.id_active ((algebra.Alg.Var 5)::nil))::
      (algebra.Alg.Var 6)::nil))
    (* first(X1_,active(X2_)) -> first(X1_,X2_) *)
   | R_xml_0_rule_40 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Var 5)::
                   (algebra.Alg.Var 6)::nil)) 
     (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Var 5)::
      (algebra.Alg.Term algebra.F.id_active 
      ((algebra.Alg.Var 6)::nil))::nil))
 .
 
 
 Definition R_xml_0_rule_as_list_0  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_terms ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_cons 
     ((algebra.Alg.Term algebra.F.id_recip ((algebra.Alg.Term 
     algebra.F.id_sqr ((algebra.Alg.Var 1)::nil))::nil))::(algebra.Alg.Term 
     algebra.F.id_terms ((algebra.Alg.Term algebra.F.id_s 
     ((algebra.Alg.Var 1)::nil))::nil))::nil))::nil)))::nil.
 
 
 Definition R_xml_0_rule_as_list_1  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_0 nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_0 
     nil)::nil)))::R_xml_0_rule_as_list_0.
 
 
 Definition R_xml_0_rule_as_list_2  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_s 
     ((algebra.Alg.Var 2)::nil))::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_s 
     ((algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term algebra.F.id_sqr 
     ((algebra.Alg.Var 2)::nil))::(algebra.Alg.Term algebra.F.id_dbl 
     ((algebra.Alg.Var 2)::nil))::nil))::nil))::nil)))::
    R_xml_0_rule_as_list_1.
 
 
 Definition R_xml_0_rule_as_list_3  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_0 nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_0 
     nil)::nil)))::R_xml_0_rule_as_list_2.
 
 
 Definition R_xml_0_rule_as_list_4  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_s 
     ((algebra.Alg.Var 2)::nil))::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_s 
     ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_dbl 
     ((algebra.Alg.Var 2)::nil))::nil))::nil))::nil)))::
    R_xml_0_rule_as_list_3.
 
 
 Definition R_xml_0_rule_as_list_5  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_add ((algebra.Alg.Term algebra.F.id_0 nil)::
     (algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 2)::nil)))::
    R_xml_0_rule_as_list_4.
 
 
 Definition R_xml_0_rule_as_list_6  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_add ((algebra.Alg.Term algebra.F.id_s 
     ((algebra.Alg.Var 2)::nil))::(algebra.Alg.Var 3)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_s 
     ((algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Var 2)::
     (algebra.Alg.Var 3)::nil))::nil))::nil)))::R_xml_0_rule_as_list_5.
 
 
 Definition R_xml_0_rule_as_list_7  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_first ((algebra.Alg.Term algebra.F.id_0 nil)::
     (algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_nil 
     nil)::nil)))::R_xml_0_rule_as_list_6.
 
 
 Definition R_xml_0_rule_as_list_8  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_first ((algebra.Alg.Term algebra.F.id_s 
     ((algebra.Alg.Var 2)::nil))::(algebra.Alg.Term algebra.F.id_cons 
     ((algebra.Alg.Var 3)::(algebra.Alg.Var 4)::nil))::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_cons 
     ((algebra.Alg.Var 3)::(algebra.Alg.Term algebra.F.id_first 
     ((algebra.Alg.Var 2)::(algebra.Alg.Var 4)::nil))::nil))::nil)))::
    R_xml_0_rule_as_list_7.
 
 
 Definition R_xml_0_rule_as_list_9  := 
   ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
     algebra.F.id_terms ((algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_terms ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 2)::nil))::nil))::nil)))::R_xml_0_rule_as_list_8.
 
 
 Definition R_xml_0_rule_as_list_10  := 
   ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_cons 
     ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Var 6)::nil))::nil)))::
    R_xml_0_rule_as_list_9.
 
 
 Definition R_xml_0_rule_as_list_11  := 
   ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
     algebra.F.id_recip ((algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_recip ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 2)::nil))::nil))::nil)))::R_xml_0_rule_as_list_10.
 
 
 Definition R_xml_0_rule_as_list_12  := 
   ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_sqr 
     ((algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 2)::nil))::nil))::nil)))::R_xml_0_rule_as_list_11.
 
 
 Definition R_xml_0_rule_as_list_13  := 
   ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_s 
     ((algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term algebra.F.id_s 
     ((algebra.Alg.Var 2)::nil))::nil)))::R_xml_0_rule_as_list_12.
 
 
 Definition R_xml_0_rule_as_list_14  := 
   ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_0 
     nil)::nil)),
    (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term algebra.F.id_0 
     nil)::nil)))::R_xml_0_rule_as_list_13.
 
 
 Definition R_xml_0_rule_as_list_15  := 
   ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_add 
     ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_add ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 6)::nil))::nil))::nil)))::R_xml_0_rule_as_list_14.
 
 
 Definition R_xml_0_rule_as_list_16  := 
   ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_dbl 
     ((algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 2)::nil))::nil))::nil)))::R_xml_0_rule_as_list_15.
 
 
 Definition R_xml_0_rule_as_list_17  := 
   ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
     algebra.F.id_first ((algebra.Alg.Var 5)::
     (algebra.Alg.Var 6)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_first ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 6)::nil))::nil))::nil)))::R_xml_0_rule_as_list_16.
 
 
 Definition R_xml_0_rule_as_list_18  := 
   ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_nil 
     nil)::nil)),
    (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_nil nil)::nil)))::R_xml_0_rule_as_list_17.
 
 
 Definition R_xml_0_rule_as_list_19  := 
   ((algebra.Alg.Term algebra.F.id_terms ((algebra.Alg.Term 
     algebra.F.id_mark ((algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_terms ((algebra.Alg.Var 2)::nil)))::
    R_xml_0_rule_as_list_18.
 
 
 Definition R_xml_0_rule_as_list_20  := 
   ((algebra.Alg.Term algebra.F.id_terms ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_terms ((algebra.Alg.Var 2)::nil)))::
    R_xml_0_rule_as_list_19.
 
 
 Definition R_xml_0_rule_as_list_21  := 
   ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Var 6)::nil)),
    (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::
     (algebra.Alg.Var 6)::nil)))::R_xml_0_rule_as_list_20.
 
 
 Definition R_xml_0_rule_as_list_22  := 
   ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::
     (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 6)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::
     (algebra.Alg.Var 6)::nil)))::R_xml_0_rule_as_list_21.
 
 
 Definition R_xml_0_rule_as_list_23  := 
   ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 5)::nil))::
     (algebra.Alg.Var 6)::nil)),
    (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::
     (algebra.Alg.Var 6)::nil)))::R_xml_0_rule_as_list_22.
 
 
 Definition R_xml_0_rule_as_list_24  := 
   ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Var 6)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::
     (algebra.Alg.Var 6)::nil)))::R_xml_0_rule_as_list_23.
 
 
 Definition R_xml_0_rule_as_list_25  := 
   ((algebra.Alg.Term algebra.F.id_recip ((algebra.Alg.Term 
     algebra.F.id_mark ((algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_recip ((algebra.Alg.Var 2)::nil)))::
    R_xml_0_rule_as_list_24.
 
 
 Definition R_xml_0_rule_as_list_26  := 
   ((algebra.Alg.Term algebra.F.id_recip ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_recip ((algebra.Alg.Var 2)::nil)))::
    R_xml_0_rule_as_list_25.
 
 
 Definition R_xml_0_rule_as_list_27  := 
   ((algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Var 2)::nil)))::
    R_xml_0_rule_as_list_26.
 
 
 Definition R_xml_0_rule_as_list_28  := 
   ((algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Var 2)::nil)))::
    R_xml_0_rule_as_list_27.
 
 
 Definition R_xml_0_rule_as_list_29  := 
   ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Var 2)::nil)))::
    R_xml_0_rule_as_list_28.
 
 
 Definition R_xml_0_rule_as_list_30  := 
   ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_active 
     ((algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Var 2)::nil)))::
    R_xml_0_rule_as_list_29.
 
 
 Definition R_xml_0_rule_as_list_31  := 
   ((algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Var 6)::nil)),
    (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Var 5)::
     (algebra.Alg.Var 6)::nil)))::R_xml_0_rule_as_list_30.
 
 
 Definition R_xml_0_rule_as_list_32  := 
   ((algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Var 5)::
     (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 6)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Var 5)::
     (algebra.Alg.Var 6)::nil)))::R_xml_0_rule_as_list_31.
 
 
 Definition R_xml_0_rule_as_list_33  := 
   ((algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 5)::nil))::
     (algebra.Alg.Var 6)::nil)),
    (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Var 5)::
     (algebra.Alg.Var 6)::nil)))::R_xml_0_rule_as_list_32.
 
 
 Definition R_xml_0_rule_as_list_34  := 
   ((algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Var 5)::
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Var 6)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Var 5)::
     (algebra.Alg.Var 6)::nil)))::R_xml_0_rule_as_list_33.
 
 
 Definition R_xml_0_rule_as_list_35  := 
   ((algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Var 2)::nil)))::
    R_xml_0_rule_as_list_34.
 
 
 Definition R_xml_0_rule_as_list_36  := 
   ((algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 2)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Var 2)::nil)))::
    R_xml_0_rule_as_list_35.
 
 
 Definition R_xml_0_rule_as_list_37  := 
   ((algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Term 
     algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::
     (algebra.Alg.Var 6)::nil)),
    (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Var 5)::
     (algebra.Alg.Var 6)::nil)))::R_xml_0_rule_as_list_36.
 
 
 Definition R_xml_0_rule_as_list_38  := 
   ((algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Var 5)::
     (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 6)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Var 5)::
     (algebra.Alg.Var 6)::nil)))::R_xml_0_rule_as_list_37.
 
 
 Definition R_xml_0_rule_as_list_39  := 
   ((algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 5)::nil))::
     (algebra.Alg.Var 6)::nil)),
    (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Var 5)::
     (algebra.Alg.Var 6)::nil)))::R_xml_0_rule_as_list_38.
 
 
 Definition R_xml_0_rule_as_list_40  := 
   ((algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Var 5)::
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Var 6)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Var 5)::
     (algebra.Alg.Var 6)::nil)))::R_xml_0_rule_as_list_39.
 
 Definition R_xml_0_rule_as_list  := R_xml_0_rule_as_list_40.
 
 
 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.or25_equiv in H|-.
   case H;clear H;intros H.
   injection H;intros ;subst;constructor 41.
   injection H;intros ;subst;constructor 40.
   injection H;intros ;subst;constructor 39.
   injection H;intros ;subst;constructor 38.
   injection H;intros ;subst;constructor 37.
   injection H;intros ;subst;constructor 36.
   injection H;intros ;subst;constructor 35.
   injection H;intros ;subst;constructor 34.
   injection H;intros ;subst;constructor 33.
   injection H;intros ;subst;constructor 32.
   injection H;intros ;subst;constructor 31.
   injection H;intros ;subst;constructor 30.
   injection H;intros ;subst;constructor 29.
   injection H;intros ;subst;constructor 28.
   injection H;intros ;subst;constructor 27.
   injection H;intros ;subst;constructor 26.
   injection H;intros ;subst;constructor 25.
   injection H;intros ;subst;constructor 24.
   injection H;intros ;subst;constructor 23.
   injection H;intros ;subst;constructor 22.
   injection H;intros ;subst;constructor 21.
   injection H;intros ;subst;constructor 20.
   injection H;intros ;subst;constructor 19.
   injection H;intros ;subst;constructor 18.
   rewrite  <- or_ext_generated.or18_equiv in H|-.
   case H;clear H;intros H.
   injection H;intros ;subst;constructor 17.
   injection H;intros ;subst;constructor 16.
   injection H;intros ;subst;constructor 15.
   injection H;intros ;subst;constructor 14.
   injection H;intros ;subst;constructor 13.
   injection H;intros ;subst;constructor 12.
   injection H;intros ;subst;constructor 11.
   injection H;intros ;subst;constructor 10.
   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_2 (x8 x9:Prop) : Prop := 
   | conj_2 : x8->x9->and_2 x8 x9
 .
 
 
 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_2 (t = (algebra.Alg.Term algebra.F.id_nil nil) ->
           t' = (algebra.Alg.Term algebra.F.id_nil nil)) 
     (t = (algebra.Alg.Term algebra.F.id_0 nil) ->
      t' = (algebra.Alg.Term algebra.F.id_0 nil)).
 Proof.
   intros t t' H.
   
   induction H as [|y IH z z_to_y] using 
   closure_extension.refl_trans_clos_ind2.
   constructor 1.
   intros H;intuition;constructor 1.
   intros H;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_nil H_id_0].
   constructor.
   
   clear H_id_0;intros 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
    ).
   
   clear H_id_nil;intros 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
    ).
 Qed.
 
 
 Lemma id_nil_is_R_xml_0_constructor :
  forall t t', 
   (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) ->
    t = (algebra.Alg.Term algebra.F.id_nil nil) ->
     t' = (algebra.Alg.Term algebra.F.id_nil nil).
 Proof.
   intros t t' H.
   destruct (are_constuctors_of_R_xml_0 H).
   assumption.
 Qed.
 
 
 Lemma id_0_is_R_xml_0_constructor :
  forall t t', 
   (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) ->
    t = (algebra.Alg.Term algebra.F.id_0 nil) ->
     t' = (algebra.Alg.Term algebra.F.id_0 nil).
 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_nil nil) |- _ =>
     let Heq := fresh "Heq" in 
      (set (Heq:=id_nil_is_R_xml_0_constructor H (refl_equal _)) in *;
        (discriminate Heq)||
        (clearbody Heq;try (subst);try (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_0 nil) |- _ =>
     let Heq := fresh "Heq" in 
      (set (Heq:=id_0_is_R_xml_0_constructor H (refl_equal _)) in *;
        (discriminate Heq)||
        (clearbody Heq;try (subst);try (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_nil nil) |- _ =>
     let Heq := fresh "Heq" in 
      (set (Heq:=id_nil_is_R_xml_0_constructor H (refl_equal _)) in *;
        (discriminate Heq)||
        (clearbody Heq;try (subst);try (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_0 nil) |- _ =>
     let Heq := fresh "Heq" in 
      (set (Heq:=id_0_is_R_xml_0_constructor H (refl_equal _)) in *;
        (discriminate Heq)||
        (clearbody Heq;try (subst);try (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_active : A ->A.
   
   Hypothesis P_id_add : A ->A ->A.
   
   Hypothesis P_id_recip : A ->A.
   
   Hypothesis P_id_mark : A ->A.
   
   Hypothesis P_id_first : A ->A ->A.
   
   Hypothesis P_id_s : A ->A.
   
   Hypothesis P_id_terms : A ->A.
   
   Hypothesis P_id_dbl : A ->A.
   
   Hypothesis P_id_sqr : A ->A.
   
   Hypothesis P_id_cons : A ->A ->A.
   
   Hypothesis P_id_nil : A.
   
   Hypothesis P_id_0 : A.
   
   Hypothesis P_id_active_monotonic :
    forall x8 x9, (A0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
   
   Hypothesis P_id_add_monotonic :
    forall x8 x10 x9 x11, 
     (A0 <= x11)/\ (x11 <= x10) ->
      (A0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
   
   Hypothesis P_id_recip_monotonic :
    forall x8 x9, (A0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
   
   Hypothesis P_id_mark_monotonic :
    forall x8 x9, (A0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
   
   Hypothesis P_id_first_monotonic :
    forall x8 x10 x9 x11, 
     (A0 <= x11)/\ (x11 <= x10) ->
      (A0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
   
   Hypothesis P_id_s_monotonic :
    forall x8 x9, (A0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
   
   Hypothesis P_id_terms_monotonic :
    forall x8 x9, (A0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
   
   Hypothesis P_id_dbl_monotonic :
    forall x8 x9, (A0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
   
   Hypothesis P_id_sqr_monotonic :
    forall x8 x9, (A0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
   
   Hypothesis P_id_cons_monotonic :
    forall x8 x10 x9 x11, 
     (A0 <= x11)/\ (x11 <= x10) ->
      (A0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
   
   Hypothesis P_id_active_bounded :
    forall x8, (A0 <= x8) ->A0 <= P_id_active x8.
   
   Hypothesis P_id_add_bounded :
    forall x8 x9, (A0 <= x8) ->(A0 <= x9) ->A0 <= P_id_add x9 x8.
   
   Hypothesis P_id_recip_bounded :
    forall x8, (A0 <= x8) ->A0 <= P_id_recip x8.
   
   Hypothesis P_id_mark_bounded : forall x8, (A0 <= x8) ->A0 <= P_id_mark x8.
   
   Hypothesis P_id_first_bounded :
    forall x8 x9, (A0 <= x8) ->(A0 <= x9) ->A0 <= P_id_first x9 x8.
   
   Hypothesis P_id_s_bounded : forall x8, (A0 <= x8) ->A0 <= P_id_s x8.
   
   Hypothesis P_id_terms_bounded :
    forall x8, (A0 <= x8) ->A0 <= P_id_terms x8.
   
   Hypothesis P_id_dbl_bounded : forall x8, (A0 <= x8) ->A0 <= P_id_dbl x8.
   
   Hypothesis P_id_sqr_bounded : forall x8, (A0 <= x8) ->A0 <= P_id_sqr x8.
   
   Hypothesis P_id_cons_bounded :
    forall x8 x9, (A0 <= x8) ->(A0 <= x9) ->A0 <= P_id_cons x9 x8.
   
   Hypothesis P_id_nil_bounded : A0 <= P_id_nil .
   
   Hypothesis P_id_0_bounded : A0 <= P_id_0 .
   
   Fixpoint measure t { struct t }  := 
     match t with
       | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
        P_id_active (measure x8)
       | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
        P_id_add (measure x9) (measure x8)
       | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
        P_id_recip (measure x8)
       | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
        P_id_mark (measure x8)
       | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
        P_id_first (measure x9) (measure x8)
       | (algebra.Alg.Term algebra.F.id_s (x8::nil)) => P_id_s (measure x8)
       | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
        P_id_terms (measure x8)
       | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
        P_id_dbl (measure x8)
       | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
        P_id_sqr (measure x8)
       | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
        P_id_cons (measure x9) (measure x8)
       | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
       | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
       | _ => A0
       end.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                    P_id_active (measure x8)
                   | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                    P_id_add (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                    P_id_recip (measure x8)
                   | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                    P_id_mark (measure x8)
                   | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                    P_id_first (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                    P_id_s (measure x8)
                   | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                    P_id_terms (measure x8)
                   | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                    P_id_dbl (measure x8)
                   | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                    P_id_sqr (measure x8)
                   | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                    P_id_cons (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                   | _ => 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_active => P_id_active
       | algebra.F.id_add => P_id_add
       | algebra.F.id_recip => P_id_recip
       | algebra.F.id_mark => P_id_mark
       | algebra.F.id_first => P_id_first
       | algebra.F.id_s => P_id_s
       | algebra.F.id_terms => P_id_terms
       | algebra.F.id_dbl => P_id_dbl
       | algebra.F.id_sqr => P_id_sqr
       | algebra.F.id_cons => P_id_cons
       | algebra.F.id_nil => P_id_nil
       | algebra.F.id_0 => P_id_0
       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_active =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_add =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_recip =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_mark =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_first =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_s =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_terms =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_dbl =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_sqr =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_cons =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_nil => match l with
                                      | nil => _
                                      | _::_ => _
                                      end
              | algebra.F.id_0 => match l with
                                    | 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_active_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_add_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_recip_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_mark_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_first_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_s_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_terms_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_dbl_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_sqr_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_cons_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_nil_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_0_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_active_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_add_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_recip_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_mark_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_first_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_s_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_terms_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_dbl_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_sqr_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_cons_monotonic;assumption.
     vm_compute in |-*;apply (Aop.(le_refl)).
     vm_compute in |-*;apply (Aop.(le_refl)).
     intros f.
     case f.
     vm_compute in |-*;intros ;apply P_id_active_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_add_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_recip_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_mark_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_first_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_s_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_terms_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_dbl_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_sqr_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_cons_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_nil_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_0_bounded;assumption.
     intros .
     do 2 (rewrite  <- same_measure in |-*).
     apply rules_monotonic;assumption.
     assumption.
   Qed.
   
   Hypothesis P_id_SQR : A ->A.
   
   Hypothesis P_id_DBL : A ->A.
   
   Hypothesis P_id_ACTIVE : A ->A.
   
   Hypothesis P_id_S : A ->A.
   
   Hypothesis P_id_CONS : A ->A ->A.
   
   Hypothesis P_id_TERMS : A ->A.
   
   Hypothesis P_id_FIRST : A ->A ->A.
   
   Hypothesis P_id_MARK : A ->A.
   
   Hypothesis P_id_ADD : A ->A ->A.
   
   Hypothesis P_id_RECIP : A ->A.
   
   Hypothesis P_id_SQR_monotonic :
    forall x8 x9, (A0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
   
   Hypothesis P_id_DBL_monotonic :
    forall x8 x9, (A0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
   
   Hypothesis P_id_ACTIVE_monotonic :
    forall x8 x9, (A0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
   
   Hypothesis P_id_S_monotonic :
    forall x8 x9, (A0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
   
   Hypothesis P_id_CONS_monotonic :
    forall x8 x10 x9 x11, 
     (A0 <= x11)/\ (x11 <= x10) ->
      (A0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
   
   Hypothesis P_id_TERMS_monotonic :
    forall x8 x9, (A0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
   
   Hypothesis P_id_FIRST_monotonic :
    forall x8 x10 x9 x11, 
     (A0 <= x11)/\ (x11 <= x10) ->
      (A0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
   
   Hypothesis P_id_MARK_monotonic :
    forall x8 x9, (A0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
   
   Hypothesis P_id_ADD_monotonic :
    forall x8 x10 x9 x11, 
     (A0 <= x11)/\ (x11 <= x10) ->
      (A0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
   
   Hypothesis P_id_RECIP_monotonic :
    forall x8 x9, (A0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
   
   Definition marked_measure t := 
     match t with
       | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
        P_id_SQR (measure x8)
       | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
        P_id_DBL (measure x8)
       | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
        P_id_ACTIVE (measure x8)
       | (algebra.Alg.Term algebra.F.id_s (x8::nil)) => P_id_S (measure x8)
       | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
        P_id_CONS (measure x9) (measure x8)
       | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
        P_id_TERMS (measure x8)
       | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
        P_id_FIRST (measure x9) (measure x8)
       | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
        P_id_MARK (measure x8)
       | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
        P_id_ADD (measure x9) (measure x8)
       | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
        P_id_RECIP (measure x8)
       | _ => 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;apply (P_id_CONS x9 x8).
     
     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;apply (P_id_SQR x8).
     
     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;apply (P_id_DBL x8).
     
     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;apply (P_id_TERMS x8).
     
     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;apply (P_id_S x8).
     
     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;apply (P_id_FIRST x9 x8).
     
     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;apply (P_id_MARK x8).
     
     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;apply (P_id_RECIP x8).
     
     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;apply (P_id_ADD x9 x8).
     
     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;apply (P_id_ACTIVE x8).
     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_active =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_add =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_recip =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_mark =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_first =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_s =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_terms =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_dbl =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_sqr =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_cons =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_nil => match l with
                                      | nil => _
                                      | _::_ => _
                                      end
              | algebra.F.id_0 => match l with
                                    | 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 .
     
     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 .
     
     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 .
     
     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 .
     
     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 .
   Qed.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                           P_id_SQR (measure x8)
                          | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                           P_id_DBL (measure x8)
                          | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                           P_id_ACTIVE (measure x8)
                          | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                           P_id_S (measure x8)
                          | (algebra.Alg.Term algebra.F.id_cons (x9::
                             x8::nil)) =>
                           P_id_CONS (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                           P_id_TERMS (measure x8)
                          | (algebra.Alg.Term algebra.F.id_first (x9::
                             x8::nil)) =>
                           P_id_FIRST (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                           P_id_MARK (measure x8)
                          | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                           P_id_ADD (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                           P_id_RECIP (measure x8)
                          | _ => 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_active_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_add_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_recip_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_mark_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_first_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_s_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_terms_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_dbl_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_sqr_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_cons_monotonic;assumption.
     vm_compute in |-*;apply (Aop.(le_refl)).
     vm_compute in |-*;apply (Aop.(le_refl)).
     clear f.
     intros f.
     case f.
     vm_compute in |-*;intros ;apply P_id_active_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_add_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_recip_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_mark_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_first_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_s_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_terms_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_dbl_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_sqr_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_cons_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_nil_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_0_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_CONS_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_SQR_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_DBL_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_TERMS_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_S_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_FIRST_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_MARK_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_RECIP_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_ADD_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_ACTIVE_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_active : Z ->Z.
   
   Hypothesis P_id_add : Z ->Z ->Z.
   
   Hypothesis P_id_recip : Z ->Z.
   
   Hypothesis P_id_mark : Z ->Z.
   
   Hypothesis P_id_first : Z ->Z ->Z.
   
   Hypothesis P_id_s : Z ->Z.
   
   Hypothesis P_id_terms : Z ->Z.
   
   Hypothesis P_id_dbl : Z ->Z.
   
   Hypothesis P_id_sqr : Z ->Z.
   
   Hypothesis P_id_cons : Z ->Z ->Z.
   
   Hypothesis P_id_nil : Z.
   
   Hypothesis P_id_0 : Z.
   
   Hypothesis P_id_active_monotonic :
    forall x8 x9, 
     (min_value <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
   
   Hypothesis P_id_add_monotonic :
    forall x8 x10 x9 x11, 
     (min_value <= x11)/\ (x11 <= x10) ->
      (min_value <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
   
   Hypothesis P_id_recip_monotonic :
    forall x8 x9, 
     (min_value <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
   
   Hypothesis P_id_mark_monotonic :
    forall x8 x9, 
     (min_value <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
   
   Hypothesis P_id_first_monotonic :
    forall x8 x10 x9 x11, 
     (min_value <= x11)/\ (x11 <= x10) ->
      (min_value <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
   
   Hypothesis P_id_s_monotonic :
    forall x8 x9, (min_value <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
   
   Hypothesis P_id_terms_monotonic :
    forall x8 x9, 
     (min_value <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
   
   Hypothesis P_id_dbl_monotonic :
    forall x8 x9, (min_value <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
   
   Hypothesis P_id_sqr_monotonic :
    forall x8 x9, (min_value <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
   
   Hypothesis P_id_cons_monotonic :
    forall x8 x10 x9 x11, 
     (min_value <= x11)/\ (x11 <= x10) ->
      (min_value <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
   
   Hypothesis P_id_active_bounded :
    forall x8, (min_value <= x8) ->min_value <= P_id_active x8.
   
   Hypothesis P_id_add_bounded :
    forall x8 x9, 
     (min_value <= x8) ->(min_value <= x9) ->min_value <= P_id_add x9 x8.
   
   Hypothesis P_id_recip_bounded :
    forall x8, (min_value <= x8) ->min_value <= P_id_recip x8.
   
   Hypothesis P_id_mark_bounded :
    forall x8, (min_value <= x8) ->min_value <= P_id_mark x8.
   
   Hypothesis P_id_first_bounded :
    forall x8 x9, 
     (min_value <= x8) ->(min_value <= x9) ->min_value <= P_id_first x9 x8.
   
   Hypothesis P_id_s_bounded :
    forall x8, (min_value <= x8) ->min_value <= P_id_s x8.
   
   Hypothesis P_id_terms_bounded :
    forall x8, (min_value <= x8) ->min_value <= P_id_terms x8.
   
   Hypothesis P_id_dbl_bounded :
    forall x8, (min_value <= x8) ->min_value <= P_id_dbl x8.
   
   Hypothesis P_id_sqr_bounded :
    forall x8, (min_value <= x8) ->min_value <= P_id_sqr x8.
   
   Hypothesis P_id_cons_bounded :
    forall x8 x9, 
     (min_value <= x8) ->(min_value <= x9) ->min_value <= P_id_cons x9 x8.
   
   Hypothesis P_id_nil_bounded : min_value <= P_id_nil .
   
   Hypothesis P_id_0_bounded : min_value <= P_id_0 .
   
   Definition measure  := 
     Interp.measure min_value P_id_active P_id_add P_id_recip P_id_mark 
      P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
      P_id_0.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                    P_id_active (measure x8)
                   | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                    P_id_add (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                    P_id_recip (measure x8)
                   | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                    P_id_mark (measure x8)
                   | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                    P_id_first (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                    P_id_s (measure x8)
                   | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                    P_id_terms (measure x8)
                   | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                    P_id_dbl (measure x8)
                   | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                    P_id_sqr (measure x8)
                   | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                    P_id_cons (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                   | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Hypothesis P_id_SQR : Z ->Z.
   
   Hypothesis P_id_DBL : Z ->Z.
   
   Hypothesis P_id_ACTIVE : Z ->Z.
   
   Hypothesis P_id_S : Z ->Z.
   
   Hypothesis P_id_CONS : Z ->Z ->Z.
   
   Hypothesis P_id_TERMS : Z ->Z.
   
   Hypothesis P_id_FIRST : Z ->Z ->Z.
   
   Hypothesis P_id_MARK : Z ->Z.
   
   Hypothesis P_id_ADD : Z ->Z ->Z.
   
   Hypothesis P_id_RECIP : Z ->Z.
   
   Hypothesis P_id_SQR_monotonic :
    forall x8 x9, (min_value <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
   
   Hypothesis P_id_DBL_monotonic :
    forall x8 x9, (min_value <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
   
   Hypothesis P_id_ACTIVE_monotonic :
    forall x8 x9, 
     (min_value <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
   
   Hypothesis P_id_S_monotonic :
    forall x8 x9, (min_value <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
   
   Hypothesis P_id_CONS_monotonic :
    forall x8 x10 x9 x11, 
     (min_value <= x11)/\ (x11 <= x10) ->
      (min_value <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
   
   Hypothesis P_id_TERMS_monotonic :
    forall x8 x9, 
     (min_value <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
   
   Hypothesis P_id_FIRST_monotonic :
    forall x8 x10 x9 x11, 
     (min_value <= x11)/\ (x11 <= x10) ->
      (min_value <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
   
   Hypothesis P_id_MARK_monotonic :
    forall x8 x9, 
     (min_value <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
   
   Hypothesis P_id_ADD_monotonic :
    forall x8 x10 x9 x11, 
     (min_value <= x11)/\ (x11 <= x10) ->
      (min_value <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
   
   Hypothesis P_id_RECIP_monotonic :
    forall x8 x9, 
     (min_value <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
   
   Definition marked_measure  := 
     Interp.marked_measure min_value P_id_active P_id_add P_id_recip 
      P_id_mark P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons 
      P_id_nil P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS 
      P_id_TERMS P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                           P_id_SQR (measure x8)
                          | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                           P_id_DBL (measure x8)
                          | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                           P_id_ACTIVE (measure x8)
                          | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                           P_id_S (measure x8)
                          | (algebra.Alg.Term algebra.F.id_cons (x9::
                             x8::nil)) =>
                           P_id_CONS (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                           P_id_TERMS (measure x8)
                          | (algebra.Alg.Term algebra.F.id_first (x9::
                             x8::nil)) =>
                           P_id_FIRST (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                           P_id_MARK (measure x8)
                          | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                           P_id_ADD (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                           P_id_RECIP (measure x8)
                          | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_SQR_monotonic;assumption.
     intros ;apply P_id_DBL_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_S_monotonic;assumption.
     intros ;apply P_id_CONS_monotonic;assumption.
     intros ;apply P_id_TERMS_monotonic;assumption.
     intros ;apply P_id_FIRST_monotonic;assumption.
     intros ;apply P_id_MARK_monotonic;assumption.
     intros ;apply P_id_ADD_monotonic;assumption.
     intros ;apply P_id_RECIP_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 := 
    (* <active(terms(N_)),mark(cons(recip(sqr(N_)),terms(s(N_))))> *)
   | DP_R_xml_0_0 :
    forall x8 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                  algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_recip 
                  ((algebra.Alg.Term algebra.F.id_sqr (x1::nil))::nil))::
                  (algebra.Alg.Term algebra.F.id_terms ((algebra.Alg.Term 
                  algebra.F.id_s (x1::nil))::nil))::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
   
    (* <active(terms(N_)),cons(recip(sqr(N_)),terms(s(N_)))> *)
   | DP_R_xml_0_1 :
    forall x8 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term 
                  algebra.F.id_recip ((algebra.Alg.Term algebra.F.id_sqr 
                  (x1::nil))::nil))::(algebra.Alg.Term algebra.F.id_terms 
                  ((algebra.Alg.Term algebra.F.id_s (x1::nil))::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(terms(N_)),recip(sqr(N_))> *)
   | DP_R_xml_0_2 :
    forall x8 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_recip ((algebra.Alg.Term 
                  algebra.F.id_sqr (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(terms(N_)),sqr(N_)> *)
   | DP_R_xml_0_3 :
    forall x8 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_sqr (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(terms(N_)),terms(s(N_))> *)
   | DP_R_xml_0_4 :
    forall x8 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_terms ((algebra.Alg.Term 
                  algebra.F.id_s (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(terms(N_)),s(N_)> *)
   | DP_R_xml_0_5 :
    forall x8 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_s (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(sqr(0)),mark(0)> *)
   | DP_R_xml_0_6 :
    forall x8, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_0 
        nil)::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                  algebra.F.id_0 nil)::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(sqr(s(X_))),mark(s(add(sqr(X_),dbl(X_))))> *)
   | DP_R_xml_0_7 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_s 
        (x2::nil))::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                  algebra.F.id_s ((algebra.Alg.Term algebra.F.id_add 
                  ((algebra.Alg.Term algebra.F.id_sqr (x2::nil))::
                  (algebra.Alg.Term algebra.F.id_dbl 
                  (x2::nil))::nil))::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(sqr(s(X_))),s(add(sqr(X_),dbl(X_)))> *)
   | DP_R_xml_0_8 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_s 
        (x2::nil))::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term 
                  algebra.F.id_add ((algebra.Alg.Term algebra.F.id_sqr 
                  (x2::nil))::(algebra.Alg.Term algebra.F.id_dbl 
                  (x2::nil))::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(sqr(s(X_))),add(sqr(X_),dbl(X_))> *)
   | DP_R_xml_0_9 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_s 
        (x2::nil))::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term 
                  algebra.F.id_sqr (x2::nil))::(algebra.Alg.Term 
                  algebra.F.id_dbl (x2::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(sqr(s(X_))),sqr(X_)> *)
   | DP_R_xml_0_10 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_s 
        (x2::nil))::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(sqr(s(X_))),dbl(X_)> *)
   | DP_R_xml_0_11 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_s 
        (x2::nil))::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(dbl(0)),mark(0)> *)
   | DP_R_xml_0_12 :
    forall x8, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_0 
        nil)::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                  algebra.F.id_0 nil)::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(dbl(s(X_))),mark(s(s(dbl(X_))))> *)
   | DP_R_xml_0_13 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_s 
        (x2::nil))::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                  algebra.F.id_s ((algebra.Alg.Term algebra.F.id_s 
                  ((algebra.Alg.Term algebra.F.id_dbl 
                  (x2::nil))::nil))::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(dbl(s(X_))),s(s(dbl(X_)))> *)
   | DP_R_xml_0_14 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_s 
        (x2::nil))::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term 
                  algebra.F.id_s ((algebra.Alg.Term algebra.F.id_dbl 
                  (x2::nil))::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(dbl(s(X_))),s(dbl(X_))> *)
   | DP_R_xml_0_15 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_s 
        (x2::nil))::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term 
                  algebra.F.id_dbl (x2::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(dbl(s(X_))),dbl(X_)> *)
   | DP_R_xml_0_16 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_s 
        (x2::nil))::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(add(0,X_)),mark(X_)> *)
   | DP_R_xml_0_17 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term algebra.F.id_0 
        nil)::x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(add(s(X_),Y_)),mark(s(add(X_,Y_)))> *)
   | DP_R_xml_0_18 :
    forall x8 x2 x3, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term algebra.F.id_s 
        (x2::nil))::x3::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                  algebra.F.id_s ((algebra.Alg.Term algebra.F.id_add (x2::
                  x3::nil))::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(add(s(X_),Y_)),s(add(X_,Y_))> *)
   | DP_R_xml_0_19 :
    forall x8 x2 x3, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term algebra.F.id_s 
        (x2::nil))::x3::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term 
                  algebra.F.id_add (x2::x3::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(add(s(X_),Y_)),add(X_,Y_)> *)
   | DP_R_xml_0_20 :
    forall x8 x2 x3, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term algebra.F.id_s 
        (x2::nil))::x3::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_add (x2::x3::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(first(0,X_)),mark(nil)> *)
   | DP_R_xml_0_21 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Term 
        algebra.F.id_0 nil)::x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                  algebra.F.id_nil nil)::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
   
    (* <active(first(s(X_),cons(Y_,Z_))),mark(cons(Y_,first(X_,Z_)))> *)
   | DP_R_xml_0_22 :
    forall x8 x4 x2 x3, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Term 
        algebra.F.id_s (x2::nil))::(algebra.Alg.Term algebra.F.id_cons (x3::
        x4::nil))::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                  algebra.F.id_cons (x3::(algebra.Alg.Term 
                  algebra.F.id_first (x2::x4::nil))::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
   
    (* <active(first(s(X_),cons(Y_,Z_))),cons(Y_,first(X_,Z_))> *)
   | DP_R_xml_0_23 :
    forall x8 x4 x2 x3, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Term 
        algebra.F.id_s (x2::nil))::(algebra.Alg.Term algebra.F.id_cons (x3::
        x4::nil))::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_cons (x3::(algebra.Alg.Term 
                  algebra.F.id_first (x2::x4::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <active(first(s(X_),cons(Y_,Z_))),first(X_,Z_)> *)
   | DP_R_xml_0_24 :
    forall x8 x4 x2 x3, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Term 
        algebra.F.id_s (x2::nil))::(algebra.Alg.Term algebra.F.id_cons (x3::
        x4::nil))::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_first (x2::x4::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x8::nil))
    (* <mark(terms(X_)),active(terms(mark(X_)))> *)
   | DP_R_xml_0_25 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_terms (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                  algebra.F.id_terms ((algebra.Alg.Term algebra.F.id_mark 
                  (x2::nil))::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(terms(X_)),terms(mark(X_))> *)
   | DP_R_xml_0_26 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_terms (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_terms ((algebra.Alg.Term 
                  algebra.F.id_mark (x2::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(terms(X_)),mark(X_)> *)
   | DP_R_xml_0_27 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_terms (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(cons(X1_,X2_)),active(cons(mark(X1_),X2_))> *)
   | DP_R_xml_0_28 :
    forall x8 x6 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                  algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_mark 
                  (x5::nil))::x6::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(cons(X1_,X2_)),cons(mark(X1_),X2_)> *)
   | DP_R_xml_0_29 :
    forall x8 x6 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term 
                  algebra.F.id_mark (x5::nil))::x6::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(cons(X1_,X2_)),mark(X1_)> *)
   | DP_R_xml_0_30 :
    forall x8 x6 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(recip(X_)),active(recip(mark(X_)))> *)
   | DP_R_xml_0_31 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_recip (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                  algebra.F.id_recip ((algebra.Alg.Term algebra.F.id_mark 
                  (x2::nil))::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(recip(X_)),recip(mark(X_))> *)
   | DP_R_xml_0_32 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_recip (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_recip ((algebra.Alg.Term 
                  algebra.F.id_mark (x2::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(recip(X_)),mark(X_)> *)
   | DP_R_xml_0_33 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_recip (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(sqr(X_)),active(sqr(mark(X_)))> *)
   | DP_R_xml_0_34 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                  algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_mark 
                  (x2::nil))::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(sqr(X_)),sqr(mark(X_))> *)
   | DP_R_xml_0_35 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term 
                  algebra.F.id_mark (x2::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(sqr(X_)),mark(X_)> *)
   | DP_R_xml_0_36 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(s(X_)),active(s(X_))> *)
   | DP_R_xml_0_37 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_s (x2::nil)) 
       x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                  algebra.F.id_s (x2::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(s(X_)),s(X_)> *)
   | DP_R_xml_0_38 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_s (x2::nil)) 
       x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_s (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(0),active(0)> *)
   | DP_R_xml_0_39 :
    forall x8, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_0 nil) 
       x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                  algebra.F.id_0 nil)::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
   
    (* <mark(add(X1_,X2_)),active(add(mark(X1_),mark(X2_)))> *)
   | DP_R_xml_0_40 :
    forall x8 x6 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                  algebra.F.id_add ((algebra.Alg.Term algebra.F.id_mark 
                  (x5::nil))::(algebra.Alg.Term algebra.F.id_mark 
                  (x6::nil))::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(add(X1_,X2_)),add(mark(X1_),mark(X2_))> *)
   | DP_R_xml_0_41 :
    forall x8 x6 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term 
                  algebra.F.id_mark (x5::nil))::(algebra.Alg.Term 
                  algebra.F.id_mark (x6::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(add(X1_,X2_)),mark(X1_)> *)
   | DP_R_xml_0_42 :
    forall x8 x6 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(add(X1_,X2_)),mark(X2_)> *)
   | DP_R_xml_0_43 :
    forall x8 x6 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x6::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(dbl(X_)),active(dbl(mark(X_)))> *)
   | DP_R_xml_0_44 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                  algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_mark 
                  (x2::nil))::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(dbl(X_)),dbl(mark(X_))> *)
   | DP_R_xml_0_45 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term 
                  algebra.F.id_mark (x2::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(dbl(X_)),mark(X_)> *)
   | DP_R_xml_0_46 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
   
    (* <mark(first(X1_,X2_)),active(first(mark(X1_),mark(X2_)))> *)
   | DP_R_xml_0_47 :
    forall x8 x6 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                  algebra.F.id_first ((algebra.Alg.Term algebra.F.id_mark 
                  (x5::nil))::(algebra.Alg.Term algebra.F.id_mark 
                  (x6::nil))::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(first(X1_,X2_)),first(mark(X1_),mark(X2_))> *)
   | DP_R_xml_0_48 :
    forall x8 x6 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Term 
                  algebra.F.id_mark (x5::nil))::(algebra.Alg.Term 
                  algebra.F.id_mark (x6::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(first(X1_,X2_)),mark(X1_)> *)
   | DP_R_xml_0_49 :
    forall x8 x6 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(first(X1_,X2_)),mark(X2_)> *)
   | DP_R_xml_0_50 :
    forall x8 x6 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x6::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <mark(nil),active(nil)> *)
   | DP_R_xml_0_51 :
    forall x8, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_nil nil) 
       x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
                  algebra.F.id_nil nil)::nil)) 
       (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    (* <terms(mark(X_)),terms(X_)> *)
   | DP_R_xml_0_52 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_terms (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_terms (x8::nil))
    (* <terms(active(X_)),terms(X_)> *)
   | DP_R_xml_0_53 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_active (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_terms (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_terms (x8::nil))
    (* <cons(mark(X1_),X2_),cons(X1_,X2_)> *)
   | DP_R_xml_0_54 :
    forall x8 x6 x9 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x5::nil)) x9) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x6 x8) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) 
        (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil))
    (* <cons(X1_,mark(X2_)),cons(X1_,X2_)> *)
   | DP_R_xml_0_55 :
    forall x8 x6 x9 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                x5 x9) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x6::nil)) x8) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) 
        (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil))
    (* <cons(active(X1_),X2_),cons(X1_,X2_)> *)
   | DP_R_xml_0_56 :
    forall x8 x6 x9 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_active (x5::nil)) x9) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x6 x8) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) 
        (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil))
    (* <cons(X1_,active(X2_)),cons(X1_,X2_)> *)
   | DP_R_xml_0_57 :
    forall x8 x6 x9 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                x5 x9) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_active (x6::nil)) x8) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) 
        (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil))
    (* <recip(mark(X_)),recip(X_)> *)
   | DP_R_xml_0_58 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_recip (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_recip (x8::nil))
    (* <recip(active(X_)),recip(X_)> *)
   | DP_R_xml_0_59 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_active (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_recip (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_recip (x8::nil))
    (* <sqr(mark(X_)),sqr(X_)> *)
   | DP_R_xml_0_60 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_sqr (x8::nil))
    (* <sqr(active(X_)),sqr(X_)> *)
   | DP_R_xml_0_61 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_active (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_sqr (x8::nil))
    (* <s(mark(X_)),s(X_)> *)
   | DP_R_xml_0_62 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_s (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_s (x8::nil))
    (* <s(active(X_)),s(X_)> *)
   | DP_R_xml_0_63 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_active (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_s (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_s (x8::nil))
    (* <add(mark(X1_),X2_),add(X1_,X2_)> *)
   | DP_R_xml_0_64 :
    forall x8 x6 x9 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x5::nil)) x9) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x6 x8) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) 
        (algebra.Alg.Term algebra.F.id_add (x9::x8::nil))
    (* <add(X1_,mark(X2_)),add(X1_,X2_)> *)
   | DP_R_xml_0_65 :
    forall x8 x6 x9 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                x5 x9) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x6::nil)) x8) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) 
        (algebra.Alg.Term algebra.F.id_add (x9::x8::nil))
    (* <add(active(X1_),X2_),add(X1_,X2_)> *)
   | DP_R_xml_0_66 :
    forall x8 x6 x9 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_active (x5::nil)) x9) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x6 x8) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) 
        (algebra.Alg.Term algebra.F.id_add (x9::x8::nil))
    (* <add(X1_,active(X2_)),add(X1_,X2_)> *)
   | DP_R_xml_0_67 :
    forall x8 x6 x9 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                x5 x9) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_active (x6::nil)) x8) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) 
        (algebra.Alg.Term algebra.F.id_add (x9::x8::nil))
    (* <dbl(mark(X_)),dbl(X_)> *)
   | DP_R_xml_0_68 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_dbl (x8::nil))
    (* <dbl(active(X_)),dbl(X_)> *)
   | DP_R_xml_0_69 :
    forall x8 x2, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_active (x2::nil)) x8) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) 
       (algebra.Alg.Term algebra.F.id_dbl (x8::nil))
    (* <first(mark(X1_),X2_),first(X1_,X2_)> *)
   | DP_R_xml_0_70 :
    forall x8 x6 x9 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x5::nil)) x9) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x6 x8) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) 
        (algebra.Alg.Term algebra.F.id_first (x9::x8::nil))
    (* <first(X1_,mark(X2_)),first(X1_,X2_)> *)
   | DP_R_xml_0_71 :
    forall x8 x6 x9 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                x5 x9) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x6::nil)) x8) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) 
        (algebra.Alg.Term algebra.F.id_first (x9::x8::nil))
    (* <first(active(X1_),X2_),first(X1_,X2_)> *)
   | DP_R_xml_0_72 :
    forall x8 x6 x9 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_active (x5::nil)) x9) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x6 x8) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) 
        (algebra.Alg.Term algebra.F.id_first (x9::x8::nil))
    (* <first(X1_,active(X2_)),first(X1_,X2_)> *)
   | DP_R_xml_0_73 :
    forall x8 x6 x9 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                x5 x9) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_active (x6::nil)) x8) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) 
        (algebra.Alg.Term algebra.F.id_first (x9::x8::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_scc_1  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <first(X1_,mark(X2_)),first(X1_,X2_)> *)
    | DP_R_xml_0_scc_1_0 :
     forall x8 x6 x9 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x5 x9) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x6::nil)) x8) ->
        DP_R_xml_0_scc_1 (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) 
         (algebra.Alg.Term algebra.F.id_first (x9::x8::nil))
     (* <first(mark(X1_),X2_),first(X1_,X2_)> *)
    | DP_R_xml_0_scc_1_1 :
     forall x8 x6 x9 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x5::nil)) x9) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x6 x8) ->
        DP_R_xml_0_scc_1 (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) 
         (algebra.Alg.Term algebra.F.id_first (x9::x8::nil))
     (* <first(active(X1_),X2_),first(X1_,X2_)> *)
    | DP_R_xml_0_scc_1_2 :
     forall x8 x6 x9 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_active (x5::nil)) x9) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x6 x8) ->
        DP_R_xml_0_scc_1 (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) 
         (algebra.Alg.Term algebra.F.id_first (x9::x8::nil))
     (* <first(X1_,active(X2_)),first(X1_,X2_)> *)
    | DP_R_xml_0_scc_1_3 :
     forall x8 x6 x9 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x5 x9) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_active (x6::nil)) x8) ->
        DP_R_xml_0_scc_1 (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) 
         (algebra.Alg.Term algebra.F.id_first (x9::x8::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_1.
   Open Scope Z_scope.
   
   Import ring_extention.
   
   Notation Local "a <= b" := (Zle a b).
   
   Notation Local "a < b" := (Zlt a b).
   
   Definition P_id_active (x8:Z) := 1 + 1* x8.
   
   Definition P_id_add (x8:Z) (x9:Z) := 2 + 2* x9.
   
   Definition P_id_recip (x8:Z) := 0.
   
   Definition P_id_mark (x8:Z) := 1 + 2* x8.
   
   Definition P_id_first (x8:Z) (x9:Z) := 0.
   
   Definition P_id_s (x8:Z) := 0.
   
   Definition P_id_terms (x8:Z) := 0.
   
   Definition P_id_dbl (x8:Z) := 0.
   
   Definition P_id_sqr (x8:Z) := 3.
   
   Definition P_id_cons (x8:Z) (x9:Z) := 2* x9.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_0  := 0.
   
   Lemma P_id_active_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
   Proof.
     intros x9 x8.
     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_add_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_recip_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
   Proof.
     intros x9 x8.
     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_mark_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
   Proof.
     intros x9 x8.
     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_first_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_s_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
   Proof.
     intros x9 x8.
     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_terms_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
   Proof.
     intros x9 x8.
     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_dbl_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
   Proof.
     intros x9 x8.
     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_sqr_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
   Proof.
     intros x9 x8.
     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_cons_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_add_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_first_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_cons_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_nil_bounded : 0 <= P_id_nil .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_0_bounded : 0 <= P_id_0 .
   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_active P_id_add P_id_recip P_id_mark P_id_first 
      P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil P_id_0.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                    P_id_active (measure x8)
                   | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                    P_id_add (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                    P_id_recip (measure x8)
                   | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                    P_id_mark (measure x8)
                   | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                    P_id_first (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                    P_id_s (measure x8)
                   | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                    P_id_terms (measure x8)
                   | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                    P_id_dbl (measure x8)
                   | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                    P_id_sqr (measure x8)
                   | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                    P_id_cons (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                   | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_SQR (x8:Z) := 0.
   
   Definition P_id_DBL (x8:Z) := 0.
   
   Definition P_id_ACTIVE (x8:Z) := 0.
   
   Definition P_id_S (x8:Z) := 0.
   
   Definition P_id_CONS (x8:Z) (x9:Z) := 0.
   
   Definition P_id_TERMS (x8:Z) := 0.
   
   Definition P_id_FIRST (x8:Z) (x9:Z) := 3* x8 + 3* x9.
   
   Definition P_id_MARK (x8:Z) := 0.
   
   Definition P_id_ADD (x8:Z) (x9:Z) := 0.
   
   Definition P_id_RECIP (x8:Z) := 0.
   
   Lemma P_id_SQR_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
   Proof.
     intros x9 x8.
     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_DBL_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
   Proof.
     intros x9 x8.
     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_ACTIVE_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
   Proof.
     intros x9 x8.
     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_S_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
   Proof.
     intros x9 x8.
     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_CONS_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_TERMS_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
   Proof.
     intros x9 x8.
     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_FIRST_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_MARK_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
   Proof.
     intros x9 x8.
     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_ADD_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_RECIP_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
   Proof.
     intros x9 x8.
     intros [H_1 H_0].
     
     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_active P_id_add P_id_recip P_id_mark 
      P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
      P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
      P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                           P_id_SQR (measure x8)
                          | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                           P_id_DBL (measure x8)
                          | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                           P_id_ACTIVE (measure x8)
                          | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                           P_id_S (measure x8)
                          | (algebra.Alg.Term algebra.F.id_cons (x9::
                             x8::nil)) =>
                           P_id_CONS (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                           P_id_TERMS (measure x8)
                          | (algebra.Alg.Term algebra.F.id_first (x9::
                             x8::nil)) =>
                           P_id_FIRST (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                           P_id_MARK (measure x8)
                          | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                           P_id_ADD (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                           P_id_RECIP (measure x8)
                          | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_SQR_monotonic;assumption.
     intros ;apply P_id_DBL_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_S_monotonic;assumption.
     intros ;apply P_id_CONS_monotonic;assumption.
     intros ;apply P_id_TERMS_monotonic;assumption.
     intros ;apply P_id_FIRST_monotonic;assumption.
     intros ;apply P_id_MARK_monotonic;assumption.
     intros ;apply P_id_ADD_monotonic;assumption.
     intros ;apply P_id_RECIP_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_1.
   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_1.
  
  Definition wf_DP_R_xml_0_scc_1  := WF_DP_R_xml_0_scc_1.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_1 :
   forall x y, (DP_R_xml_0_scc_1 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_1).
    intros x' _ Hrec y h.
    
    inversion h;clear h;subst;
     constructor;intros _y _h;inversion _h;clear _h;subst;
      (eapply Hrec;econstructor eassumption)||
      ((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_1.
  Qed.
  
  
  Inductive DP_R_xml_0_scc_2  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <dbl(active(X_)),dbl(X_)> *)
    | DP_R_xml_0_scc_2_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_active (x2::nil)) x8) ->
       DP_R_xml_0_scc_2 (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_dbl (x8::nil))
     (* <dbl(mark(X_)),dbl(X_)> *)
    | DP_R_xml_0_scc_2_1 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x2::nil)) x8) ->
       DP_R_xml_0_scc_2 (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_dbl (x8::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_2.
   Inductive DP_R_xml_0_scc_2_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <dbl(active(X_)),dbl(X_)> *)
     | DP_R_xml_0_scc_2_large_0 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_active (x2::nil)) x8) ->
        DP_R_xml_0_scc_2_large (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) 
         (algebra.Alg.Term algebra.F.id_dbl (x8::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_2_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <dbl(mark(X_)),dbl(X_)> *)
     | DP_R_xml_0_scc_2_strict_0 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x2::nil)) x8) ->
        DP_R_xml_0_scc_2_strict (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) 
         (algebra.Alg.Term algebra.F.id_dbl (x8::nil))
   .
   
   
   Module WF_DP_R_xml_0_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_active (x8:Z) := 1 + 2* x8.
    
    Definition P_id_add (x8:Z) (x9:Z) := 0.
    
    Definition P_id_recip (x8:Z) := 0.
    
    Definition P_id_mark (x8:Z) := 1.
    
    Definition P_id_first (x8:Z) (x9:Z) := 0.
    
    Definition P_id_s (x8:Z) := 0.
    
    Definition P_id_terms (x8:Z) := 0.
    
    Definition P_id_dbl (x8:Z) := 0.
    
    Definition P_id_sqr (x8:Z) := 0.
    
    Definition P_id_cons (x8:Z) (x9:Z) := 0.
    
    Definition P_id_nil  := 0.
    
    Definition P_id_0  := 0.
    
    Lemma P_id_active_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
    Proof.
      intros x9 x8.
      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_add_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_recip_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
    Proof.
      intros x9 x8.
      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_mark_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
    Proof.
      intros x9 x8.
      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_first_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_s_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
    Proof.
      intros x9 x8.
      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_terms_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
    Proof.
      intros x9 x8.
      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_dbl_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
    Proof.
      intros x9 x8.
      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_sqr_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
    Proof.
      intros x9 x8.
      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_cons_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_add_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_first_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_cons_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_nil_bounded : 0 <= P_id_nil .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_0_bounded : 0 <= P_id_0 .
    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_active P_id_add P_id_recip P_id_mark P_id_first 
       P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil P_id_0.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                     P_id_active (measure x8)
                    | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                     P_id_add (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                     P_id_recip (measure x8)
                    | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                     P_id_mark (measure x8)
                    | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                     P_id_first (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                     P_id_s (measure x8)
                    | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                     P_id_terms (measure x8)
                    | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                     P_id_dbl (measure x8)
                    | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                     P_id_sqr (measure x8)
                    | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                     P_id_cons (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                    | _ => 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_active_monotonic;assumption.
      intros ;apply P_id_add_monotonic;assumption.
      intros ;apply P_id_recip_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_first_monotonic;assumption.
      intros ;apply P_id_s_monotonic;assumption.
      intros ;apply P_id_terms_monotonic;assumption.
      intros ;apply P_id_dbl_monotonic;assumption.
      intros ;apply P_id_sqr_monotonic;assumption.
      intros ;apply P_id_cons_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_add_bounded;assumption.
      intros ;apply P_id_recip_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_first_bounded;assumption.
      intros ;apply P_id_s_bounded;assumption.
      intros ;apply P_id_terms_bounded;assumption.
      intros ;apply P_id_dbl_bounded;assumption.
      intros ;apply P_id_sqr_bounded;assumption.
      intros ;apply P_id_cons_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_0_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_SQR (x8:Z) := 0.
    
    Definition P_id_DBL (x8:Z) := 1* x8.
    
    Definition P_id_ACTIVE (x8:Z) := 0.
    
    Definition P_id_S (x8:Z) := 0.
    
    Definition P_id_CONS (x8:Z) (x9:Z) := 0.
    
    Definition P_id_TERMS (x8:Z) := 0.
    
    Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
    
    Definition P_id_MARK (x8:Z) := 0.
    
    Definition P_id_ADD (x8:Z) (x9:Z) := 0.
    
    Definition P_id_RECIP (x8:Z) := 0.
    
    Lemma P_id_SQR_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
    Proof.
      intros x9 x8.
      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_DBL_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
    Proof.
      intros x9 x8.
      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_ACTIVE_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
    Proof.
      intros x9 x8.
      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_S_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
    Proof.
      intros x9 x8.
      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_CONS_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_TERMS_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
    Proof.
      intros x9 x8.
      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_FIRST_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_MARK_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
    Proof.
      intros x9 x8.
      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_ADD_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_RECIP_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
    Proof.
      intros x9 x8.
      intros [H_1 H_0].
      
      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_active P_id_add P_id_recip P_id_mark 
       P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
       P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
       P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                            P_id_SQR (measure x8)
                           | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                            P_id_DBL (measure x8)
                           | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                            P_id_ACTIVE (measure x8)
                           | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                            P_id_S (measure x8)
                           | (algebra.Alg.Term algebra.F.id_cons (x9::
                              x8::nil)) =>
                            P_id_CONS (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                            P_id_TERMS (measure x8)
                           | (algebra.Alg.Term algebra.F.id_first (x9::
                              x8::nil)) =>
                            P_id_FIRST (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                            P_id_MARK (measure x8)
                           | (algebra.Alg.Term algebra.F.id_add (x9::
                              x8::nil)) =>
                            P_id_ADD (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                            P_id_RECIP (measure x8)
                           | _ => 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_active_monotonic;assumption.
      intros ;apply P_id_add_monotonic;assumption.
      intros ;apply P_id_recip_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_first_monotonic;assumption.
      intros ;apply P_id_s_monotonic;assumption.
      intros ;apply P_id_terms_monotonic;assumption.
      intros ;apply P_id_dbl_monotonic;assumption.
      intros ;apply P_id_sqr_monotonic;assumption.
      intros ;apply P_id_cons_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_add_bounded;assumption.
      intros ;apply P_id_recip_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_first_bounded;assumption.
      intros ;apply P_id_s_bounded;assumption.
      intros ;apply P_id_terms_bounded;assumption.
      intros ;apply P_id_dbl_bounded;assumption.
      intros ;apply P_id_sqr_bounded;assumption.
      intros ;apply P_id_cons_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_0_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_SQR_monotonic;assumption.
      intros ;apply P_id_DBL_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_S_monotonic;assumption.
      intros ;apply P_id_CONS_monotonic;assumption.
      intros ;apply P_id_TERMS_monotonic;assumption.
      intros ;apply P_id_FIRST_monotonic;assumption.
      intros ;apply P_id_MARK_monotonic;assumption.
      intros ;apply P_id_ADD_monotonic;assumption.
      intros ;apply P_id_RECIP_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_2.DP_R_xml_0_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_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_active (x8:Z) := 1* x8.
   
   Definition P_id_add (x8:Z) (x9:Z) := 2 + 2* x9.
   
   Definition P_id_recip (x8:Z) := 2.
   
   Definition P_id_mark (x8:Z) := 1 + 2* x8.
   
   Definition P_id_first (x8:Z) (x9:Z) := 1 + 1* x9.
   
   Definition P_id_s (x8:Z) := 0.
   
   Definition P_id_terms (x8:Z) := 2.
   
   Definition P_id_dbl (x8:Z) := 2 + 2* x8.
   
   Definition P_id_sqr (x8:Z) := 1.
   
   Definition P_id_cons (x8:Z) (x9:Z) := 0.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_0  := 0.
   
   Lemma P_id_active_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
   Proof.
     intros x9 x8.
     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_add_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_recip_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
   Proof.
     intros x9 x8.
     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_mark_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
   Proof.
     intros x9 x8.
     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_first_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_s_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
   Proof.
     intros x9 x8.
     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_terms_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
   Proof.
     intros x9 x8.
     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_dbl_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
   Proof.
     intros x9 x8.
     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_sqr_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
   Proof.
     intros x9 x8.
     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_cons_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_add_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_first_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_cons_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_nil_bounded : 0 <= P_id_nil .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_0_bounded : 0 <= P_id_0 .
   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_active P_id_add P_id_recip P_id_mark P_id_first 
      P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil P_id_0.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                    P_id_active (measure x8)
                   | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                    P_id_add (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                    P_id_recip (measure x8)
                   | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                    P_id_mark (measure x8)
                   | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                    P_id_first (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                    P_id_s (measure x8)
                   | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                    P_id_terms (measure x8)
                   | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                    P_id_dbl (measure x8)
                   | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                    P_id_sqr (measure x8)
                   | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                    P_id_cons (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                   | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_SQR (x8:Z) := 0.
   
   Definition P_id_DBL (x8:Z) := 1* x8.
   
   Definition P_id_ACTIVE (x8:Z) := 0.
   
   Definition P_id_S (x8:Z) := 0.
   
   Definition P_id_CONS (x8:Z) (x9:Z) := 0.
   
   Definition P_id_TERMS (x8:Z) := 0.
   
   Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
   
   Definition P_id_MARK (x8:Z) := 0.
   
   Definition P_id_ADD (x8:Z) (x9:Z) := 0.
   
   Definition P_id_RECIP (x8:Z) := 0.
   
   Lemma P_id_SQR_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
   Proof.
     intros x9 x8.
     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_DBL_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
   Proof.
     intros x9 x8.
     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_ACTIVE_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
   Proof.
     intros x9 x8.
     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_S_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
   Proof.
     intros x9 x8.
     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_CONS_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_TERMS_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
   Proof.
     intros x9 x8.
     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_FIRST_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_MARK_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
   Proof.
     intros x9 x8.
     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_ADD_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_RECIP_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
   Proof.
     intros x9 x8.
     intros [H_1 H_0].
     
     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_active P_id_add P_id_recip P_id_mark 
      P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
      P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
      P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                           P_id_SQR (measure x8)
                          | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                           P_id_DBL (measure x8)
                          | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                           P_id_ACTIVE (measure x8)
                          | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                           P_id_S (measure x8)
                          | (algebra.Alg.Term algebra.F.id_cons (x9::
                             x8::nil)) =>
                           P_id_CONS (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                           P_id_TERMS (measure x8)
                          | (algebra.Alg.Term algebra.F.id_first (x9::
                             x8::nil)) =>
                           P_id_FIRST (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                           P_id_MARK (measure x8)
                          | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                           P_id_ADD (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                           P_id_RECIP (measure x8)
                          | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_SQR_monotonic;assumption.
     intros ;apply P_id_DBL_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_S_monotonic;assumption.
     intros ;apply P_id_CONS_monotonic;assumption.
     intros ;apply P_id_TERMS_monotonic;assumption.
     intros ;apply P_id_FIRST_monotonic;assumption.
     intros ;apply P_id_MARK_monotonic;assumption.
     intros ;apply P_id_ADD_monotonic;assumption.
     intros ;apply P_id_RECIP_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_2_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_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_2_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_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_2_large  := WF_DP_R_xml_0_scc_2_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_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_2_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_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_2_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_2_large.
   Qed.
  End WF_DP_R_xml_0_scc_2.
  
  Definition wf_DP_R_xml_0_scc_2  := WF_DP_R_xml_0_scc_2.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_2 :
   forall x y, (DP_R_xml_0_scc_2 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_2).
    intros x' _ Hrec y h.
    
    inversion h;clear h;subst;
     constructor;intros _y _h;inversion _h;clear _h;subst;
      (eapply Hrec;econstructor eassumption)||
      ((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_2.
  Qed.
  
  
  Inductive DP_R_xml_0_scc_3  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <add(X1_,mark(X2_)),add(X1_,X2_)> *)
    | DP_R_xml_0_scc_3_0 :
     forall x8 x6 x9 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x5 x9) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x6::nil)) x8) ->
        DP_R_xml_0_scc_3 (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) 
         (algebra.Alg.Term algebra.F.id_add (x9::x8::nil))
     (* <add(mark(X1_),X2_),add(X1_,X2_)> *)
    | DP_R_xml_0_scc_3_1 :
     forall x8 x6 x9 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x5::nil)) x9) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x6 x8) ->
        DP_R_xml_0_scc_3 (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) 
         (algebra.Alg.Term algebra.F.id_add (x9::x8::nil))
     (* <add(active(X1_),X2_),add(X1_,X2_)> *)
    | DP_R_xml_0_scc_3_2 :
     forall x8 x6 x9 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_active (x5::nil)) x9) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x6 x8) ->
        DP_R_xml_0_scc_3 (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) 
         (algebra.Alg.Term algebra.F.id_add (x9::x8::nil))
     (* <add(X1_,active(X2_)),add(X1_,X2_)> *)
    | DP_R_xml_0_scc_3_3 :
     forall x8 x6 x9 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x5 x9) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_active (x6::nil)) x8) ->
        DP_R_xml_0_scc_3 (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) 
         (algebra.Alg.Term algebra.F.id_add (x9::x8::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_3.
   Open Scope Z_scope.
   
   Import ring_extention.
   
   Notation Local "a <= b" := (Zle a b).
   
   Notation Local "a < b" := (Zlt a b).
   
   Definition P_id_active (x8:Z) := 1 + 1* x8.
   
   Definition P_id_add (x8:Z) (x9:Z) := 2 + 2* x9.
   
   Definition P_id_recip (x8:Z) := 0.
   
   Definition P_id_mark (x8:Z) := 1 + 2* x8.
   
   Definition P_id_first (x8:Z) (x9:Z) := 0.
   
   Definition P_id_s (x8:Z) := 0.
   
   Definition P_id_terms (x8:Z) := 0.
   
   Definition P_id_dbl (x8:Z) := 0.
   
   Definition P_id_sqr (x8:Z) := 3.
   
   Definition P_id_cons (x8:Z) (x9:Z) := 2* x9.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_0  := 0.
   
   Lemma P_id_active_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
   Proof.
     intros x9 x8.
     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_add_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_recip_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
   Proof.
     intros x9 x8.
     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_mark_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
   Proof.
     intros x9 x8.
     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_first_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_s_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
   Proof.
     intros x9 x8.
     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_terms_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
   Proof.
     intros x9 x8.
     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_dbl_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
   Proof.
     intros x9 x8.
     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_sqr_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
   Proof.
     intros x9 x8.
     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_cons_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_add_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_first_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_cons_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_nil_bounded : 0 <= P_id_nil .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_0_bounded : 0 <= P_id_0 .
   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_active P_id_add P_id_recip P_id_mark P_id_first 
      P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil P_id_0.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                    P_id_active (measure x8)
                   | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                    P_id_add (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                    P_id_recip (measure x8)
                   | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                    P_id_mark (measure x8)
                   | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                    P_id_first (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                    P_id_s (measure x8)
                   | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                    P_id_terms (measure x8)
                   | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                    P_id_dbl (measure x8)
                   | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                    P_id_sqr (measure x8)
                   | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                    P_id_cons (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                   | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_SQR (x8:Z) := 0.
   
   Definition P_id_DBL (x8:Z) := 0.
   
   Definition P_id_ACTIVE (x8:Z) := 0.
   
   Definition P_id_S (x8:Z) := 0.
   
   Definition P_id_CONS (x8:Z) (x9:Z) := 0.
   
   Definition P_id_TERMS (x8:Z) := 0.
   
   Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
   
   Definition P_id_MARK (x8:Z) := 0.
   
   Definition P_id_ADD (x8:Z) (x9:Z) := 3* x8 + 3* x9.
   
   Definition P_id_RECIP (x8:Z) := 0.
   
   Lemma P_id_SQR_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
   Proof.
     intros x9 x8.
     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_DBL_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
   Proof.
     intros x9 x8.
     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_ACTIVE_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
   Proof.
     intros x9 x8.
     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_S_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
   Proof.
     intros x9 x8.
     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_CONS_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_TERMS_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
   Proof.
     intros x9 x8.
     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_FIRST_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_MARK_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
   Proof.
     intros x9 x8.
     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_ADD_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_RECIP_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
   Proof.
     intros x9 x8.
     intros [H_1 H_0].
     
     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_active P_id_add P_id_recip P_id_mark 
      P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
      P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
      P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                           P_id_SQR (measure x8)
                          | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                           P_id_DBL (measure x8)
                          | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                           P_id_ACTIVE (measure x8)
                          | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                           P_id_S (measure x8)
                          | (algebra.Alg.Term algebra.F.id_cons (x9::
                             x8::nil)) =>
                           P_id_CONS (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                           P_id_TERMS (measure x8)
                          | (algebra.Alg.Term algebra.F.id_first (x9::
                             x8::nil)) =>
                           P_id_FIRST (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                           P_id_MARK (measure x8)
                          | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                           P_id_ADD (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                           P_id_RECIP (measure x8)
                          | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_SQR_monotonic;assumption.
     intros ;apply P_id_DBL_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_S_monotonic;assumption.
     intros ;apply P_id_CONS_monotonic;assumption.
     intros ;apply P_id_TERMS_monotonic;assumption.
     intros ;apply P_id_FIRST_monotonic;assumption.
     intros ;apply P_id_MARK_monotonic;assumption.
     intros ;apply P_id_ADD_monotonic;assumption.
     intros ;apply P_id_RECIP_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_3.
   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_3.
  
  Definition wf_DP_R_xml_0_scc_3  := WF_DP_R_xml_0_scc_3.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_3 :
   forall x y, (DP_R_xml_0_scc_3 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_3).
    intros x' _ Hrec y h.
    
    inversion h;clear h;subst;
     constructor;intros _y _h;inversion _h;clear _h;subst;
      (eapply Hrec;econstructor eassumption)||
      ((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_3.
  Qed.
  
  
  Inductive DP_R_xml_0_scc_4  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <s(active(X_)),s(X_)> *)
    | DP_R_xml_0_scc_4_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_active (x2::nil)) x8) ->
       DP_R_xml_0_scc_4 (algebra.Alg.Term algebra.F.id_s (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_s (x8::nil))
     (* <s(mark(X_)),s(X_)> *)
    | DP_R_xml_0_scc_4_1 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x2::nil)) x8) ->
       DP_R_xml_0_scc_4 (algebra.Alg.Term algebra.F.id_s (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_s (x8::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_4.
   Inductive DP_R_xml_0_scc_4_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <s(active(X_)),s(X_)> *)
     | DP_R_xml_0_scc_4_large_0 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_active (x2::nil)) x8) ->
        DP_R_xml_0_scc_4_large (algebra.Alg.Term algebra.F.id_s (x2::nil)) 
         (algebra.Alg.Term algebra.F.id_s (x8::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_4_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <s(mark(X_)),s(X_)> *)
     | DP_R_xml_0_scc_4_strict_0 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x2::nil)) x8) ->
        DP_R_xml_0_scc_4_strict (algebra.Alg.Term algebra.F.id_s (x2::nil)) 
         (algebra.Alg.Term algebra.F.id_s (x8::nil))
   .
   
   
   Module 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_active (x8:Z) := 1 + 2* x8.
    
    Definition P_id_add (x8:Z) (x9:Z) := 0.
    
    Definition P_id_recip (x8:Z) := 0.
    
    Definition P_id_mark (x8:Z) := 1.
    
    Definition P_id_first (x8:Z) (x9:Z) := 0.
    
    Definition P_id_s (x8:Z) := 0.
    
    Definition P_id_terms (x8:Z) := 0.
    
    Definition P_id_dbl (x8:Z) := 0.
    
    Definition P_id_sqr (x8:Z) := 0.
    
    Definition P_id_cons (x8:Z) (x9:Z) := 0.
    
    Definition P_id_nil  := 0.
    
    Definition P_id_0  := 0.
    
    Lemma P_id_active_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
    Proof.
      intros x9 x8.
      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_add_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_recip_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
    Proof.
      intros x9 x8.
      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_mark_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
    Proof.
      intros x9 x8.
      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_first_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_s_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
    Proof.
      intros x9 x8.
      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_terms_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
    Proof.
      intros x9 x8.
      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_dbl_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
    Proof.
      intros x9 x8.
      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_sqr_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
    Proof.
      intros x9 x8.
      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_cons_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_add_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_first_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_cons_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_nil_bounded : 0 <= P_id_nil .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_0_bounded : 0 <= P_id_0 .
    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_active P_id_add P_id_recip P_id_mark P_id_first 
       P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil P_id_0.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                     P_id_active (measure x8)
                    | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                     P_id_add (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                     P_id_recip (measure x8)
                    | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                     P_id_mark (measure x8)
                    | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                     P_id_first (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                     P_id_s (measure x8)
                    | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                     P_id_terms (measure x8)
                    | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                     P_id_dbl (measure x8)
                    | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                     P_id_sqr (measure x8)
                    | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                     P_id_cons (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                    | _ => 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_active_monotonic;assumption.
      intros ;apply P_id_add_monotonic;assumption.
      intros ;apply P_id_recip_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_first_monotonic;assumption.
      intros ;apply P_id_s_monotonic;assumption.
      intros ;apply P_id_terms_monotonic;assumption.
      intros ;apply P_id_dbl_monotonic;assumption.
      intros ;apply P_id_sqr_monotonic;assumption.
      intros ;apply P_id_cons_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_add_bounded;assumption.
      intros ;apply P_id_recip_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_first_bounded;assumption.
      intros ;apply P_id_s_bounded;assumption.
      intros ;apply P_id_terms_bounded;assumption.
      intros ;apply P_id_dbl_bounded;assumption.
      intros ;apply P_id_sqr_bounded;assumption.
      intros ;apply P_id_cons_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_0_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_SQR (x8:Z) := 0.
    
    Definition P_id_DBL (x8:Z) := 0.
    
    Definition P_id_ACTIVE (x8:Z) := 0.
    
    Definition P_id_S (x8:Z) := 1* x8.
    
    Definition P_id_CONS (x8:Z) (x9:Z) := 0.
    
    Definition P_id_TERMS (x8:Z) := 0.
    
    Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
    
    Definition P_id_MARK (x8:Z) := 0.
    
    Definition P_id_ADD (x8:Z) (x9:Z) := 0.
    
    Definition P_id_RECIP (x8:Z) := 0.
    
    Lemma P_id_SQR_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
    Proof.
      intros x9 x8.
      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_DBL_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
    Proof.
      intros x9 x8.
      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_ACTIVE_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
    Proof.
      intros x9 x8.
      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_S_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
    Proof.
      intros x9 x8.
      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_CONS_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_TERMS_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
    Proof.
      intros x9 x8.
      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_FIRST_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_MARK_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
    Proof.
      intros x9 x8.
      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_ADD_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_RECIP_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
    Proof.
      intros x9 x8.
      intros [H_1 H_0].
      
      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_active P_id_add P_id_recip P_id_mark 
       P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
       P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
       P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                            P_id_SQR (measure x8)
                           | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                            P_id_DBL (measure x8)
                           | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                            P_id_ACTIVE (measure x8)
                           | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                            P_id_S (measure x8)
                           | (algebra.Alg.Term algebra.F.id_cons (x9::
                              x8::nil)) =>
                            P_id_CONS (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                            P_id_TERMS (measure x8)
                           | (algebra.Alg.Term algebra.F.id_first (x9::
                              x8::nil)) =>
                            P_id_FIRST (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                            P_id_MARK (measure x8)
                           | (algebra.Alg.Term algebra.F.id_add (x9::
                              x8::nil)) =>
                            P_id_ADD (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                            P_id_RECIP (measure x8)
                           | _ => 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_active_monotonic;assumption.
      intros ;apply P_id_add_monotonic;assumption.
      intros ;apply P_id_recip_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_first_monotonic;assumption.
      intros ;apply P_id_s_monotonic;assumption.
      intros ;apply P_id_terms_monotonic;assumption.
      intros ;apply P_id_dbl_monotonic;assumption.
      intros ;apply P_id_sqr_monotonic;assumption.
      intros ;apply P_id_cons_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_add_bounded;assumption.
      intros ;apply P_id_recip_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_first_bounded;assumption.
      intros ;apply P_id_s_bounded;assumption.
      intros ;apply P_id_terms_bounded;assumption.
      intros ;apply P_id_dbl_bounded;assumption.
      intros ;apply P_id_sqr_bounded;assumption.
      intros ;apply P_id_cons_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_0_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_SQR_monotonic;assumption.
      intros ;apply P_id_DBL_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_S_monotonic;assumption.
      intros ;apply P_id_CONS_monotonic;assumption.
      intros ;apply P_id_TERMS_monotonic;assumption.
      intros ;apply P_id_FIRST_monotonic;assumption.
      intros ;apply P_id_MARK_monotonic;assumption.
      intros ;apply P_id_ADD_monotonic;assumption.
      intros ;apply P_id_RECIP_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.DP_R_xml_0_scc_4_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.
   
   Open Scope Z_scope.
   
   Import ring_extention.
   
   Notation Local "a <= b" := (Zle a b).
   
   Notation Local "a < b" := (Zlt a b).
   
   Definition P_id_active (x8:Z) := 1* x8.
   
   Definition P_id_add (x8:Z) (x9:Z) := 2 + 2* x9.
   
   Definition P_id_recip (x8:Z) := 2.
   
   Definition P_id_mark (x8:Z) := 1 + 2* x8.
   
   Definition P_id_first (x8:Z) (x9:Z) := 1 + 1* x9.
   
   Definition P_id_s (x8:Z) := 0.
   
   Definition P_id_terms (x8:Z) := 2.
   
   Definition P_id_dbl (x8:Z) := 2 + 2* x8.
   
   Definition P_id_sqr (x8:Z) := 1.
   
   Definition P_id_cons (x8:Z) (x9:Z) := 0.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_0  := 0.
   
   Lemma P_id_active_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
   Proof.
     intros x9 x8.
     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_add_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_recip_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
   Proof.
     intros x9 x8.
     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_mark_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
   Proof.
     intros x9 x8.
     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_first_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_s_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
   Proof.
     intros x9 x8.
     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_terms_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
   Proof.
     intros x9 x8.
     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_dbl_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
   Proof.
     intros x9 x8.
     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_sqr_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
   Proof.
     intros x9 x8.
     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_cons_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_add_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_first_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_cons_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_nil_bounded : 0 <= P_id_nil .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_0_bounded : 0 <= P_id_0 .
   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_active P_id_add P_id_recip P_id_mark P_id_first 
      P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil P_id_0.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                    P_id_active (measure x8)
                   | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                    P_id_add (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                    P_id_recip (measure x8)
                   | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                    P_id_mark (measure x8)
                   | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                    P_id_first (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                    P_id_s (measure x8)
                   | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                    P_id_terms (measure x8)
                   | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                    P_id_dbl (measure x8)
                   | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                    P_id_sqr (measure x8)
                   | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                    P_id_cons (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                   | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_SQR (x8:Z) := 0.
   
   Definition P_id_DBL (x8:Z) := 0.
   
   Definition P_id_ACTIVE (x8:Z) := 0.
   
   Definition P_id_S (x8:Z) := 1* x8.
   
   Definition P_id_CONS (x8:Z) (x9:Z) := 0.
   
   Definition P_id_TERMS (x8:Z) := 0.
   
   Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
   
   Definition P_id_MARK (x8:Z) := 0.
   
   Definition P_id_ADD (x8:Z) (x9:Z) := 0.
   
   Definition P_id_RECIP (x8:Z) := 0.
   
   Lemma P_id_SQR_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
   Proof.
     intros x9 x8.
     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_DBL_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
   Proof.
     intros x9 x8.
     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_ACTIVE_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
   Proof.
     intros x9 x8.
     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_S_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
   Proof.
     intros x9 x8.
     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_CONS_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_TERMS_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
   Proof.
     intros x9 x8.
     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_FIRST_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_MARK_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
   Proof.
     intros x9 x8.
     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_ADD_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_RECIP_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
   Proof.
     intros x9 x8.
     intros [H_1 H_0].
     
     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_active P_id_add P_id_recip P_id_mark 
      P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
      P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
      P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                           P_id_SQR (measure x8)
                          | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                           P_id_DBL (measure x8)
                          | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                           P_id_ACTIVE (measure x8)
                          | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                           P_id_S (measure x8)
                          | (algebra.Alg.Term algebra.F.id_cons (x9::
                             x8::nil)) =>
                           P_id_CONS (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                           P_id_TERMS (measure x8)
                          | (algebra.Alg.Term algebra.F.id_first (x9::
                             x8::nil)) =>
                           P_id_FIRST (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                           P_id_MARK (measure x8)
                          | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                           P_id_ADD (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                           P_id_RECIP (measure x8)
                          | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_SQR_monotonic;assumption.
     intros ;apply P_id_DBL_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_S_monotonic;assumption.
     intros ;apply P_id_CONS_monotonic;assumption.
     intros ;apply P_id_TERMS_monotonic;assumption.
     intros ;apply P_id_FIRST_monotonic;assumption.
     intros ;apply P_id_MARK_monotonic;assumption.
     intros ;apply P_id_ADD_monotonic;assumption.
     intros ;apply P_id_RECIP_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)||
      ((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.
  
  
  Inductive DP_R_xml_0_scc_5  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <sqr(active(X_)),sqr(X_)> *)
    | DP_R_xml_0_scc_5_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_active (x2::nil)) x8) ->
       DP_R_xml_0_scc_5 (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_sqr (x8::nil))
     (* <sqr(mark(X_)),sqr(X_)> *)
    | DP_R_xml_0_scc_5_1 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x2::nil)) x8) ->
       DP_R_xml_0_scc_5 (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_sqr (x8::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_5.
   Inductive DP_R_xml_0_scc_5_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <sqr(active(X_)),sqr(X_)> *)
     | DP_R_xml_0_scc_5_large_0 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_active (x2::nil)) x8) ->
        DP_R_xml_0_scc_5_large (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) 
         (algebra.Alg.Term algebra.F.id_sqr (x8::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_5_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <sqr(mark(X_)),sqr(X_)> *)
     | DP_R_xml_0_scc_5_strict_0 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x2::nil)) x8) ->
        DP_R_xml_0_scc_5_strict (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) 
         (algebra.Alg.Term algebra.F.id_sqr (x8::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_5_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_active (x8:Z) := 1 + 2* x8.
    
    Definition P_id_add (x8:Z) (x9:Z) := 0.
    
    Definition P_id_recip (x8:Z) := 0.
    
    Definition P_id_mark (x8:Z) := 1.
    
    Definition P_id_first (x8:Z) (x9:Z) := 0.
    
    Definition P_id_s (x8:Z) := 0.
    
    Definition P_id_terms (x8:Z) := 0.
    
    Definition P_id_dbl (x8:Z) := 0.
    
    Definition P_id_sqr (x8:Z) := 0.
    
    Definition P_id_cons (x8:Z) (x9:Z) := 0.
    
    Definition P_id_nil  := 0.
    
    Definition P_id_0  := 0.
    
    Lemma P_id_active_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
    Proof.
      intros x9 x8.
      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_add_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_recip_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
    Proof.
      intros x9 x8.
      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_mark_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
    Proof.
      intros x9 x8.
      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_first_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_s_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
    Proof.
      intros x9 x8.
      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_terms_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
    Proof.
      intros x9 x8.
      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_dbl_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
    Proof.
      intros x9 x8.
      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_sqr_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
    Proof.
      intros x9 x8.
      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_cons_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_add_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_first_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_cons_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_nil_bounded : 0 <= P_id_nil .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_0_bounded : 0 <= P_id_0 .
    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_active P_id_add P_id_recip P_id_mark P_id_first 
       P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil P_id_0.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                     P_id_active (measure x8)
                    | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                     P_id_add (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                     P_id_recip (measure x8)
                    | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                     P_id_mark (measure x8)
                    | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                     P_id_first (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                     P_id_s (measure x8)
                    | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                     P_id_terms (measure x8)
                    | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                     P_id_dbl (measure x8)
                    | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                     P_id_sqr (measure x8)
                    | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                     P_id_cons (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                    | _ => 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_active_monotonic;assumption.
      intros ;apply P_id_add_monotonic;assumption.
      intros ;apply P_id_recip_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_first_monotonic;assumption.
      intros ;apply P_id_s_monotonic;assumption.
      intros ;apply P_id_terms_monotonic;assumption.
      intros ;apply P_id_dbl_monotonic;assumption.
      intros ;apply P_id_sqr_monotonic;assumption.
      intros ;apply P_id_cons_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_add_bounded;assumption.
      intros ;apply P_id_recip_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_first_bounded;assumption.
      intros ;apply P_id_s_bounded;assumption.
      intros ;apply P_id_terms_bounded;assumption.
      intros ;apply P_id_dbl_bounded;assumption.
      intros ;apply P_id_sqr_bounded;assumption.
      intros ;apply P_id_cons_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_0_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_SQR (x8:Z) := 1* x8.
    
    Definition P_id_DBL (x8:Z) := 0.
    
    Definition P_id_ACTIVE (x8:Z) := 0.
    
    Definition P_id_S (x8:Z) := 0.
    
    Definition P_id_CONS (x8:Z) (x9:Z) := 0.
    
    Definition P_id_TERMS (x8:Z) := 0.
    
    Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
    
    Definition P_id_MARK (x8:Z) := 0.
    
    Definition P_id_ADD (x8:Z) (x9:Z) := 0.
    
    Definition P_id_RECIP (x8:Z) := 0.
    
    Lemma P_id_SQR_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
    Proof.
      intros x9 x8.
      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_DBL_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
    Proof.
      intros x9 x8.
      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_ACTIVE_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
    Proof.
      intros x9 x8.
      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_S_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
    Proof.
      intros x9 x8.
      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_CONS_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_TERMS_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
    Proof.
      intros x9 x8.
      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_FIRST_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_MARK_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
    Proof.
      intros x9 x8.
      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_ADD_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_RECIP_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
    Proof.
      intros x9 x8.
      intros [H_1 H_0].
      
      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_active P_id_add P_id_recip P_id_mark 
       P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
       P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
       P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                            P_id_SQR (measure x8)
                           | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                            P_id_DBL (measure x8)
                           | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                            P_id_ACTIVE (measure x8)
                           | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                            P_id_S (measure x8)
                           | (algebra.Alg.Term algebra.F.id_cons (x9::
                              x8::nil)) =>
                            P_id_CONS (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                            P_id_TERMS (measure x8)
                           | (algebra.Alg.Term algebra.F.id_first (x9::
                              x8::nil)) =>
                            P_id_FIRST (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                            P_id_MARK (measure x8)
                           | (algebra.Alg.Term algebra.F.id_add (x9::
                              x8::nil)) =>
                            P_id_ADD (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                            P_id_RECIP (measure x8)
                           | _ => 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_active_monotonic;assumption.
      intros ;apply P_id_add_monotonic;assumption.
      intros ;apply P_id_recip_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_first_monotonic;assumption.
      intros ;apply P_id_s_monotonic;assumption.
      intros ;apply P_id_terms_monotonic;assumption.
      intros ;apply P_id_dbl_monotonic;assumption.
      intros ;apply P_id_sqr_monotonic;assumption.
      intros ;apply P_id_cons_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_add_bounded;assumption.
      intros ;apply P_id_recip_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_first_bounded;assumption.
      intros ;apply P_id_s_bounded;assumption.
      intros ;apply P_id_terms_bounded;assumption.
      intros ;apply P_id_dbl_bounded;assumption.
      intros ;apply P_id_sqr_bounded;assumption.
      intros ;apply P_id_cons_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_0_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_SQR_monotonic;assumption.
      intros ;apply P_id_DBL_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_S_monotonic;assumption.
      intros ;apply P_id_CONS_monotonic;assumption.
      intros ;apply P_id_TERMS_monotonic;assumption.
      intros ;apply P_id_FIRST_monotonic;assumption.
      intros ;apply P_id_MARK_monotonic;assumption.
      intros ;apply P_id_ADD_monotonic;assumption.
      intros ;apply P_id_RECIP_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_5.DP_R_xml_0_scc_5_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_5_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_active (x8:Z) := 1* x8.
   
   Definition P_id_add (x8:Z) (x9:Z) := 2 + 2* x9.
   
   Definition P_id_recip (x8:Z) := 2.
   
   Definition P_id_mark (x8:Z) := 1 + 2* x8.
   
   Definition P_id_first (x8:Z) (x9:Z) := 1 + 1* x9.
   
   Definition P_id_s (x8:Z) := 0.
   
   Definition P_id_terms (x8:Z) := 2.
   
   Definition P_id_dbl (x8:Z) := 2 + 2* x8.
   
   Definition P_id_sqr (x8:Z) := 1.
   
   Definition P_id_cons (x8:Z) (x9:Z) := 0.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_0  := 0.
   
   Lemma P_id_active_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
   Proof.
     intros x9 x8.
     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_add_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_recip_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
   Proof.
     intros x9 x8.
     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_mark_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
   Proof.
     intros x9 x8.
     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_first_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_s_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
   Proof.
     intros x9 x8.
     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_terms_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
   Proof.
     intros x9 x8.
     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_dbl_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
   Proof.
     intros x9 x8.
     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_sqr_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
   Proof.
     intros x9 x8.
     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_cons_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_add_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_first_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_cons_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_nil_bounded : 0 <= P_id_nil .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_0_bounded : 0 <= P_id_0 .
   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_active P_id_add P_id_recip P_id_mark P_id_first 
      P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil P_id_0.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                    P_id_active (measure x8)
                   | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                    P_id_add (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                    P_id_recip (measure x8)
                   | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                    P_id_mark (measure x8)
                   | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                    P_id_first (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                    P_id_s (measure x8)
                   | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                    P_id_terms (measure x8)
                   | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                    P_id_dbl (measure x8)
                   | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                    P_id_sqr (measure x8)
                   | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                    P_id_cons (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                   | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_SQR (x8:Z) := 1* x8.
   
   Definition P_id_DBL (x8:Z) := 0.
   
   Definition P_id_ACTIVE (x8:Z) := 0.
   
   Definition P_id_S (x8:Z) := 0.
   
   Definition P_id_CONS (x8:Z) (x9:Z) := 0.
   
   Definition P_id_TERMS (x8:Z) := 0.
   
   Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
   
   Definition P_id_MARK (x8:Z) := 0.
   
   Definition P_id_ADD (x8:Z) (x9:Z) := 0.
   
   Definition P_id_RECIP (x8:Z) := 0.
   
   Lemma P_id_SQR_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
   Proof.
     intros x9 x8.
     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_DBL_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
   Proof.
     intros x9 x8.
     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_ACTIVE_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
   Proof.
     intros x9 x8.
     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_S_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
   Proof.
     intros x9 x8.
     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_CONS_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_TERMS_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
   Proof.
     intros x9 x8.
     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_FIRST_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_MARK_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
   Proof.
     intros x9 x8.
     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_ADD_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_RECIP_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
   Proof.
     intros x9 x8.
     intros [H_1 H_0].
     
     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_active P_id_add P_id_recip P_id_mark 
      P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
      P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
      P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                           P_id_SQR (measure x8)
                          | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                           P_id_DBL (measure x8)
                          | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                           P_id_ACTIVE (measure x8)
                          | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                           P_id_S (measure x8)
                          | (algebra.Alg.Term algebra.F.id_cons (x9::
                             x8::nil)) =>
                           P_id_CONS (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                           P_id_TERMS (measure x8)
                          | (algebra.Alg.Term algebra.F.id_first (x9::
                             x8::nil)) =>
                           P_id_FIRST (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                           P_id_MARK (measure x8)
                          | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                           P_id_ADD (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                           P_id_RECIP (measure x8)
                          | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_SQR_monotonic;assumption.
     intros ;apply P_id_DBL_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_S_monotonic;assumption.
     intros ;apply P_id_CONS_monotonic;assumption.
     intros ;apply P_id_TERMS_monotonic;assumption.
     intros ;apply P_id_FIRST_monotonic;assumption.
     intros ;apply P_id_MARK_monotonic;assumption.
     intros ;apply P_id_ADD_monotonic;assumption.
     intros ;apply P_id_RECIP_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_5_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_5_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_5_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_5_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_5_large  := WF_DP_R_xml_0_scc_5_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_5.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_5_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_5_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_5_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_5_large.
   Qed.
  End WF_DP_R_xml_0_scc_5.
  
  Definition wf_DP_R_xml_0_scc_5  := WF_DP_R_xml_0_scc_5.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_5 :
   forall x y, (DP_R_xml_0_scc_5 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_5).
    intros x' _ Hrec y h.
    
    inversion h;clear h;subst;
     constructor;intros _y _h;inversion _h;clear _h;subst;
      (eapply Hrec;econstructor eassumption)||
      ((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_5.
  Qed.
  
  
  Inductive DP_R_xml_0_scc_6  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <recip(active(X_)),recip(X_)> *)
    | DP_R_xml_0_scc_6_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_active (x2::nil)) x8) ->
       DP_R_xml_0_scc_6 (algebra.Alg.Term algebra.F.id_recip (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_recip (x8::nil))
     (* <recip(mark(X_)),recip(X_)> *)
    | DP_R_xml_0_scc_6_1 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x2::nil)) x8) ->
       DP_R_xml_0_scc_6 (algebra.Alg.Term algebra.F.id_recip (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_recip (x8::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_6.
   Inductive DP_R_xml_0_scc_6_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <recip(active(X_)),recip(X_)> *)
     | DP_R_xml_0_scc_6_large_0 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_active (x2::nil)) x8) ->
        DP_R_xml_0_scc_6_large (algebra.Alg.Term algebra.F.id_recip 
                                (x2::nil)) 
         (algebra.Alg.Term algebra.F.id_recip (x8::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_6_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <recip(mark(X_)),recip(X_)> *)
     | DP_R_xml_0_scc_6_strict_0 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x2::nil)) x8) ->
        DP_R_xml_0_scc_6_strict (algebra.Alg.Term algebra.F.id_recip 
                                 (x2::nil)) 
         (algebra.Alg.Term algebra.F.id_recip (x8::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_6_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_active (x8:Z) := 1 + 2* x8.
    
    Definition P_id_add (x8:Z) (x9:Z) := 0.
    
    Definition P_id_recip (x8:Z) := 0.
    
    Definition P_id_mark (x8:Z) := 1.
    
    Definition P_id_first (x8:Z) (x9:Z) := 0.
    
    Definition P_id_s (x8:Z) := 0.
    
    Definition P_id_terms (x8:Z) := 0.
    
    Definition P_id_dbl (x8:Z) := 0.
    
    Definition P_id_sqr (x8:Z) := 0.
    
    Definition P_id_cons (x8:Z) (x9:Z) := 0.
    
    Definition P_id_nil  := 0.
    
    Definition P_id_0  := 0.
    
    Lemma P_id_active_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
    Proof.
      intros x9 x8.
      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_add_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_recip_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
    Proof.
      intros x9 x8.
      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_mark_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
    Proof.
      intros x9 x8.
      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_first_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_s_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
    Proof.
      intros x9 x8.
      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_terms_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
    Proof.
      intros x9 x8.
      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_dbl_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
    Proof.
      intros x9 x8.
      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_sqr_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
    Proof.
      intros x9 x8.
      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_cons_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_add_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_first_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_cons_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_nil_bounded : 0 <= P_id_nil .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_0_bounded : 0 <= P_id_0 .
    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_active P_id_add P_id_recip P_id_mark P_id_first 
       P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil P_id_0.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                     P_id_active (measure x8)
                    | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                     P_id_add (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                     P_id_recip (measure x8)
                    | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                     P_id_mark (measure x8)
                    | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                     P_id_first (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                     P_id_s (measure x8)
                    | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                     P_id_terms (measure x8)
                    | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                     P_id_dbl (measure x8)
                    | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                     P_id_sqr (measure x8)
                    | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                     P_id_cons (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                    | _ => 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_active_monotonic;assumption.
      intros ;apply P_id_add_monotonic;assumption.
      intros ;apply P_id_recip_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_first_monotonic;assumption.
      intros ;apply P_id_s_monotonic;assumption.
      intros ;apply P_id_terms_monotonic;assumption.
      intros ;apply P_id_dbl_monotonic;assumption.
      intros ;apply P_id_sqr_monotonic;assumption.
      intros ;apply P_id_cons_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_add_bounded;assumption.
      intros ;apply P_id_recip_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_first_bounded;assumption.
      intros ;apply P_id_s_bounded;assumption.
      intros ;apply P_id_terms_bounded;assumption.
      intros ;apply P_id_dbl_bounded;assumption.
      intros ;apply P_id_sqr_bounded;assumption.
      intros ;apply P_id_cons_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_0_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_SQR (x8:Z) := 0.
    
    Definition P_id_DBL (x8:Z) := 0.
    
    Definition P_id_ACTIVE (x8:Z) := 0.
    
    Definition P_id_S (x8:Z) := 0.
    
    Definition P_id_CONS (x8:Z) (x9:Z) := 0.
    
    Definition P_id_TERMS (x8:Z) := 0.
    
    Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
    
    Definition P_id_MARK (x8:Z) := 0.
    
    Definition P_id_ADD (x8:Z) (x9:Z) := 0.
    
    Definition P_id_RECIP (x8:Z) := 1* x8.
    
    Lemma P_id_SQR_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
    Proof.
      intros x9 x8.
      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_DBL_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
    Proof.
      intros x9 x8.
      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_ACTIVE_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
    Proof.
      intros x9 x8.
      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_S_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
    Proof.
      intros x9 x8.
      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_CONS_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_TERMS_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
    Proof.
      intros x9 x8.
      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_FIRST_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_MARK_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
    Proof.
      intros x9 x8.
      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_ADD_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_RECIP_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
    Proof.
      intros x9 x8.
      intros [H_1 H_0].
      
      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_active P_id_add P_id_recip P_id_mark 
       P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
       P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
       P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                            P_id_SQR (measure x8)
                           | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                            P_id_DBL (measure x8)
                           | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                            P_id_ACTIVE (measure x8)
                           | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                            P_id_S (measure x8)
                           | (algebra.Alg.Term algebra.F.id_cons (x9::
                              x8::nil)) =>
                            P_id_CONS (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                            P_id_TERMS (measure x8)
                           | (algebra.Alg.Term algebra.F.id_first (x9::
                              x8::nil)) =>
                            P_id_FIRST (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                            P_id_MARK (measure x8)
                           | (algebra.Alg.Term algebra.F.id_add (x9::
                              x8::nil)) =>
                            P_id_ADD (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                            P_id_RECIP (measure x8)
                           | _ => 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_active_monotonic;assumption.
      intros ;apply P_id_add_monotonic;assumption.
      intros ;apply P_id_recip_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_first_monotonic;assumption.
      intros ;apply P_id_s_monotonic;assumption.
      intros ;apply P_id_terms_monotonic;assumption.
      intros ;apply P_id_dbl_monotonic;assumption.
      intros ;apply P_id_sqr_monotonic;assumption.
      intros ;apply P_id_cons_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_add_bounded;assumption.
      intros ;apply P_id_recip_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_first_bounded;assumption.
      intros ;apply P_id_s_bounded;assumption.
      intros ;apply P_id_terms_bounded;assumption.
      intros ;apply P_id_dbl_bounded;assumption.
      intros ;apply P_id_sqr_bounded;assumption.
      intros ;apply P_id_cons_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_0_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_SQR_monotonic;assumption.
      intros ;apply P_id_DBL_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_S_monotonic;assumption.
      intros ;apply P_id_CONS_monotonic;assumption.
      intros ;apply P_id_TERMS_monotonic;assumption.
      intros ;apply P_id_FIRST_monotonic;assumption.
      intros ;apply P_id_MARK_monotonic;assumption.
      intros ;apply P_id_ADD_monotonic;assumption.
      intros ;apply P_id_RECIP_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_6.DP_R_xml_0_scc_6_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_6_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_active (x8:Z) := 1* x8.
   
   Definition P_id_add (x8:Z) (x9:Z) := 2 + 2* x9.
   
   Definition P_id_recip (x8:Z) := 2.
   
   Definition P_id_mark (x8:Z) := 1 + 2* x8.
   
   Definition P_id_first (x8:Z) (x9:Z) := 1 + 1* x9.
   
   Definition P_id_s (x8:Z) := 0.
   
   Definition P_id_terms (x8:Z) := 2.
   
   Definition P_id_dbl (x8:Z) := 2 + 2* x8.
   
   Definition P_id_sqr (x8:Z) := 1.
   
   Definition P_id_cons (x8:Z) (x9:Z) := 0.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_0  := 0.
   
   Lemma P_id_active_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
   Proof.
     intros x9 x8.
     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_add_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_recip_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
   Proof.
     intros x9 x8.
     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_mark_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
   Proof.
     intros x9 x8.
     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_first_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_s_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
   Proof.
     intros x9 x8.
     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_terms_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
   Proof.
     intros x9 x8.
     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_dbl_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
   Proof.
     intros x9 x8.
     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_sqr_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
   Proof.
     intros x9 x8.
     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_cons_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_add_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_first_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_cons_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_nil_bounded : 0 <= P_id_nil .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_0_bounded : 0 <= P_id_0 .
   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_active P_id_add P_id_recip P_id_mark P_id_first 
      P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil P_id_0.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                    P_id_active (measure x8)
                   | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                    P_id_add (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                    P_id_recip (measure x8)
                   | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                    P_id_mark (measure x8)
                   | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                    P_id_first (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                    P_id_s (measure x8)
                   | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                    P_id_terms (measure x8)
                   | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                    P_id_dbl (measure x8)
                   | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                    P_id_sqr (measure x8)
                   | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                    P_id_cons (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                   | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_SQR (x8:Z) := 0.
   
   Definition P_id_DBL (x8:Z) := 0.
   
   Definition P_id_ACTIVE (x8:Z) := 0.
   
   Definition P_id_S (x8:Z) := 0.
   
   Definition P_id_CONS (x8:Z) (x9:Z) := 0.
   
   Definition P_id_TERMS (x8:Z) := 0.
   
   Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
   
   Definition P_id_MARK (x8:Z) := 0.
   
   Definition P_id_ADD (x8:Z) (x9:Z) := 0.
   
   Definition P_id_RECIP (x8:Z) := 1* x8.
   
   Lemma P_id_SQR_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
   Proof.
     intros x9 x8.
     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_DBL_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
   Proof.
     intros x9 x8.
     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_ACTIVE_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
   Proof.
     intros x9 x8.
     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_S_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
   Proof.
     intros x9 x8.
     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_CONS_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_TERMS_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
   Proof.
     intros x9 x8.
     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_FIRST_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_MARK_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
   Proof.
     intros x9 x8.
     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_ADD_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_RECIP_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
   Proof.
     intros x9 x8.
     intros [H_1 H_0].
     
     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_active P_id_add P_id_recip P_id_mark 
      P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
      P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
      P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                           P_id_SQR (measure x8)
                          | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                           P_id_DBL (measure x8)
                          | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                           P_id_ACTIVE (measure x8)
                          | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                           P_id_S (measure x8)
                          | (algebra.Alg.Term algebra.F.id_cons (x9::
                             x8::nil)) =>
                           P_id_CONS (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                           P_id_TERMS (measure x8)
                          | (algebra.Alg.Term algebra.F.id_first (x9::
                             x8::nil)) =>
                           P_id_FIRST (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                           P_id_MARK (measure x8)
                          | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                           P_id_ADD (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                           P_id_RECIP (measure x8)
                          | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_SQR_monotonic;assumption.
     intros ;apply P_id_DBL_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_S_monotonic;assumption.
     intros ;apply P_id_CONS_monotonic;assumption.
     intros ;apply P_id_TERMS_monotonic;assumption.
     intros ;apply P_id_FIRST_monotonic;assumption.
     intros ;apply P_id_MARK_monotonic;assumption.
     intros ;apply P_id_ADD_monotonic;assumption.
     intros ;apply P_id_RECIP_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_6_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_6_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_6_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_6_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_6_large  := WF_DP_R_xml_0_scc_6_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_6.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_6_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_6_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_6_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_6_large.
   Qed.
  End WF_DP_R_xml_0_scc_6.
  
  Definition wf_DP_R_xml_0_scc_6  := WF_DP_R_xml_0_scc_6.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_6 :
   forall x y, (DP_R_xml_0_scc_6 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_6).
    intros x' _ Hrec y h.
    
    inversion h;clear h;subst;
     constructor;intros _y _h;inversion _h;clear _h;subst;
      (eapply Hrec;econstructor eassumption)||
      ((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_6.
  Qed.
  
  
  Inductive DP_R_xml_0_scc_7  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <cons(X1_,mark(X2_)),cons(X1_,X2_)> *)
    | DP_R_xml_0_scc_7_0 :
     forall x8 x6 x9 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x5 x9) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x6::nil)) x8) ->
        DP_R_xml_0_scc_7 (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) 
         (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil))
     (* <cons(mark(X1_),X2_),cons(X1_,X2_)> *)
    | DP_R_xml_0_scc_7_1 :
     forall x8 x6 x9 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x5::nil)) x9) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x6 x8) ->
        DP_R_xml_0_scc_7 (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) 
         (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil))
     (* <cons(active(X1_),X2_),cons(X1_,X2_)> *)
    | DP_R_xml_0_scc_7_2 :
     forall x8 x6 x9 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_active (x5::nil)) x9) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x6 x8) ->
        DP_R_xml_0_scc_7 (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) 
         (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil))
     (* <cons(X1_,active(X2_)),cons(X1_,X2_)> *)
    | DP_R_xml_0_scc_7_3 :
     forall x8 x6 x9 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x5 x9) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_active (x6::nil)) x8) ->
        DP_R_xml_0_scc_7 (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) 
         (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_7.
   Open Scope Z_scope.
   
   Import ring_extention.
   
   Notation Local "a <= b" := (Zle a b).
   
   Notation Local "a < b" := (Zlt a b).
   
   Definition P_id_active (x8:Z) := 1 + 1* x8.
   
   Definition P_id_add (x8:Z) (x9:Z) := 2 + 2* x9.
   
   Definition P_id_recip (x8:Z) := 0.
   
   Definition P_id_mark (x8:Z) := 1 + 2* x8.
   
   Definition P_id_first (x8:Z) (x9:Z) := 0.
   
   Definition P_id_s (x8:Z) := 0.
   
   Definition P_id_terms (x8:Z) := 0.
   
   Definition P_id_dbl (x8:Z) := 0.
   
   Definition P_id_sqr (x8:Z) := 3.
   
   Definition P_id_cons (x8:Z) (x9:Z) := 2* x9.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_0  := 0.
   
   Lemma P_id_active_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
   Proof.
     intros x9 x8.
     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_add_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_recip_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
   Proof.
     intros x9 x8.
     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_mark_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
   Proof.
     intros x9 x8.
     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_first_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_s_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
   Proof.
     intros x9 x8.
     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_terms_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
   Proof.
     intros x9 x8.
     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_dbl_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
   Proof.
     intros x9 x8.
     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_sqr_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
   Proof.
     intros x9 x8.
     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_cons_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_add_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_first_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_cons_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_nil_bounded : 0 <= P_id_nil .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_0_bounded : 0 <= P_id_0 .
   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_active P_id_add P_id_recip P_id_mark P_id_first 
      P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil P_id_0.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                    P_id_active (measure x8)
                   | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                    P_id_add (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                    P_id_recip (measure x8)
                   | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                    P_id_mark (measure x8)
                   | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                    P_id_first (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                    P_id_s (measure x8)
                   | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                    P_id_terms (measure x8)
                   | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                    P_id_dbl (measure x8)
                   | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                    P_id_sqr (measure x8)
                   | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                    P_id_cons (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                   | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_SQR (x8:Z) := 0.
   
   Definition P_id_DBL (x8:Z) := 0.
   
   Definition P_id_ACTIVE (x8:Z) := 0.
   
   Definition P_id_S (x8:Z) := 0.
   
   Definition P_id_CONS (x8:Z) (x9:Z) := 3* x8 + 3* x9.
   
   Definition P_id_TERMS (x8:Z) := 0.
   
   Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
   
   Definition P_id_MARK (x8:Z) := 0.
   
   Definition P_id_ADD (x8:Z) (x9:Z) := 0.
   
   Definition P_id_RECIP (x8:Z) := 0.
   
   Lemma P_id_SQR_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
   Proof.
     intros x9 x8.
     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_DBL_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
   Proof.
     intros x9 x8.
     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_ACTIVE_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
   Proof.
     intros x9 x8.
     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_S_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
   Proof.
     intros x9 x8.
     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_CONS_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_TERMS_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
   Proof.
     intros x9 x8.
     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_FIRST_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_MARK_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
   Proof.
     intros x9 x8.
     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_ADD_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_RECIP_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
   Proof.
     intros x9 x8.
     intros [H_1 H_0].
     
     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_active P_id_add P_id_recip P_id_mark 
      P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
      P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
      P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                           P_id_SQR (measure x8)
                          | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                           P_id_DBL (measure x8)
                          | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                           P_id_ACTIVE (measure x8)
                          | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                           P_id_S (measure x8)
                          | (algebra.Alg.Term algebra.F.id_cons (x9::
                             x8::nil)) =>
                           P_id_CONS (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                           P_id_TERMS (measure x8)
                          | (algebra.Alg.Term algebra.F.id_first (x9::
                             x8::nil)) =>
                           P_id_FIRST (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                           P_id_MARK (measure x8)
                          | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                           P_id_ADD (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                           P_id_RECIP (measure x8)
                          | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_SQR_monotonic;assumption.
     intros ;apply P_id_DBL_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_S_monotonic;assumption.
     intros ;apply P_id_CONS_monotonic;assumption.
     intros ;apply P_id_TERMS_monotonic;assumption.
     intros ;apply P_id_FIRST_monotonic;assumption.
     intros ;apply P_id_MARK_monotonic;assumption.
     intros ;apply P_id_ADD_monotonic;assumption.
     intros ;apply P_id_RECIP_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_7.
   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_7.
  
  Definition wf_DP_R_xml_0_scc_7  := WF_DP_R_xml_0_scc_7.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_7 :
   forall x y, (DP_R_xml_0_scc_7 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_7).
    intros x' _ Hrec y h.
    
    inversion h;clear h;subst;
     constructor;intros _y _h;inversion _h;clear _h;subst;
      (eapply Hrec;econstructor eassumption)||
      ((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_7.
  Qed.
  
  
  Inductive DP_R_xml_0_scc_8  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <terms(active(X_)),terms(X_)> *)
    | DP_R_xml_0_scc_8_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_active (x2::nil)) x8) ->
       DP_R_xml_0_scc_8 (algebra.Alg.Term algebra.F.id_terms (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_terms (x8::nil))
     (* <terms(mark(X_)),terms(X_)> *)
    | DP_R_xml_0_scc_8_1 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x2::nil)) x8) ->
       DP_R_xml_0_scc_8 (algebra.Alg.Term algebra.F.id_terms (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_terms (x8::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_8.
   Inductive DP_R_xml_0_scc_8_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <terms(active(X_)),terms(X_)> *)
     | DP_R_xml_0_scc_8_large_0 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_active (x2::nil)) x8) ->
        DP_R_xml_0_scc_8_large (algebra.Alg.Term algebra.F.id_terms 
                                (x2::nil)) 
         (algebra.Alg.Term algebra.F.id_terms (x8::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_8_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <terms(mark(X_)),terms(X_)> *)
     | DP_R_xml_0_scc_8_strict_0 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x2::nil)) x8) ->
        DP_R_xml_0_scc_8_strict (algebra.Alg.Term algebra.F.id_terms 
                                 (x2::nil)) 
         (algebra.Alg.Term algebra.F.id_terms (x8::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_8_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_active (x8:Z) := 1 + 2* x8.
    
    Definition P_id_add (x8:Z) (x9:Z) := 0.
    
    Definition P_id_recip (x8:Z) := 0.
    
    Definition P_id_mark (x8:Z) := 1.
    
    Definition P_id_first (x8:Z) (x9:Z) := 0.
    
    Definition P_id_s (x8:Z) := 0.
    
    Definition P_id_terms (x8:Z) := 0.
    
    Definition P_id_dbl (x8:Z) := 0.
    
    Definition P_id_sqr (x8:Z) := 0.
    
    Definition P_id_cons (x8:Z) (x9:Z) := 0.
    
    Definition P_id_nil  := 0.
    
    Definition P_id_0  := 0.
    
    Lemma P_id_active_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
    Proof.
      intros x9 x8.
      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_add_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_recip_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
    Proof.
      intros x9 x8.
      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_mark_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
    Proof.
      intros x9 x8.
      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_first_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_s_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
    Proof.
      intros x9 x8.
      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_terms_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
    Proof.
      intros x9 x8.
      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_dbl_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
    Proof.
      intros x9 x8.
      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_sqr_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
    Proof.
      intros x9 x8.
      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_cons_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_add_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_first_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_cons_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_nil_bounded : 0 <= P_id_nil .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_0_bounded : 0 <= P_id_0 .
    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_active P_id_add P_id_recip P_id_mark P_id_first 
       P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil P_id_0.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                     P_id_active (measure x8)
                    | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                     P_id_add (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                     P_id_recip (measure x8)
                    | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                     P_id_mark (measure x8)
                    | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                     P_id_first (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                     P_id_s (measure x8)
                    | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                     P_id_terms (measure x8)
                    | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                     P_id_dbl (measure x8)
                    | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                     P_id_sqr (measure x8)
                    | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                     P_id_cons (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                    | _ => 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_active_monotonic;assumption.
      intros ;apply P_id_add_monotonic;assumption.
      intros ;apply P_id_recip_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_first_monotonic;assumption.
      intros ;apply P_id_s_monotonic;assumption.
      intros ;apply P_id_terms_monotonic;assumption.
      intros ;apply P_id_dbl_monotonic;assumption.
      intros ;apply P_id_sqr_monotonic;assumption.
      intros ;apply P_id_cons_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_add_bounded;assumption.
      intros ;apply P_id_recip_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_first_bounded;assumption.
      intros ;apply P_id_s_bounded;assumption.
      intros ;apply P_id_terms_bounded;assumption.
      intros ;apply P_id_dbl_bounded;assumption.
      intros ;apply P_id_sqr_bounded;assumption.
      intros ;apply P_id_cons_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_0_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_SQR (x8:Z) := 0.
    
    Definition P_id_DBL (x8:Z) := 0.
    
    Definition P_id_ACTIVE (x8:Z) := 0.
    
    Definition P_id_S (x8:Z) := 0.
    
    Definition P_id_CONS (x8:Z) (x9:Z) := 0.
    
    Definition P_id_TERMS (x8:Z) := 1* x8.
    
    Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
    
    Definition P_id_MARK (x8:Z) := 0.
    
    Definition P_id_ADD (x8:Z) (x9:Z) := 0.
    
    Definition P_id_RECIP (x8:Z) := 0.
    
    Lemma P_id_SQR_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
    Proof.
      intros x9 x8.
      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_DBL_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
    Proof.
      intros x9 x8.
      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_ACTIVE_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
    Proof.
      intros x9 x8.
      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_S_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
    Proof.
      intros x9 x8.
      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_CONS_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_TERMS_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
    Proof.
      intros x9 x8.
      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_FIRST_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_MARK_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
    Proof.
      intros x9 x8.
      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_ADD_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_RECIP_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
    Proof.
      intros x9 x8.
      intros [H_1 H_0].
      
      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_active P_id_add P_id_recip P_id_mark 
       P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
       P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
       P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                            P_id_SQR (measure x8)
                           | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                            P_id_DBL (measure x8)
                           | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                            P_id_ACTIVE (measure x8)
                           | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                            P_id_S (measure x8)
                           | (algebra.Alg.Term algebra.F.id_cons (x9::
                              x8::nil)) =>
                            P_id_CONS (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                            P_id_TERMS (measure x8)
                           | (algebra.Alg.Term algebra.F.id_first (x9::
                              x8::nil)) =>
                            P_id_FIRST (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                            P_id_MARK (measure x8)
                           | (algebra.Alg.Term algebra.F.id_add (x9::
                              x8::nil)) =>
                            P_id_ADD (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                            P_id_RECIP (measure x8)
                           | _ => 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_active_monotonic;assumption.
      intros ;apply P_id_add_monotonic;assumption.
      intros ;apply P_id_recip_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_first_monotonic;assumption.
      intros ;apply P_id_s_monotonic;assumption.
      intros ;apply P_id_terms_monotonic;assumption.
      intros ;apply P_id_dbl_monotonic;assumption.
      intros ;apply P_id_sqr_monotonic;assumption.
      intros ;apply P_id_cons_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_add_bounded;assumption.
      intros ;apply P_id_recip_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_first_bounded;assumption.
      intros ;apply P_id_s_bounded;assumption.
      intros ;apply P_id_terms_bounded;assumption.
      intros ;apply P_id_dbl_bounded;assumption.
      intros ;apply P_id_sqr_bounded;assumption.
      intros ;apply P_id_cons_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_0_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_SQR_monotonic;assumption.
      intros ;apply P_id_DBL_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_S_monotonic;assumption.
      intros ;apply P_id_CONS_monotonic;assumption.
      intros ;apply P_id_TERMS_monotonic;assumption.
      intros ;apply P_id_FIRST_monotonic;assumption.
      intros ;apply P_id_MARK_monotonic;assumption.
      intros ;apply P_id_ADD_monotonic;assumption.
      intros ;apply P_id_RECIP_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_8.DP_R_xml_0_scc_8_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_8_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_active (x8:Z) := 1* x8.
   
   Definition P_id_add (x8:Z) (x9:Z) := 2 + 2* x9.
   
   Definition P_id_recip (x8:Z) := 2.
   
   Definition P_id_mark (x8:Z) := 1 + 2* x8.
   
   Definition P_id_first (x8:Z) (x9:Z) := 1 + 1* x9.
   
   Definition P_id_s (x8:Z) := 0.
   
   Definition P_id_terms (x8:Z) := 2.
   
   Definition P_id_dbl (x8:Z) := 2 + 2* x8.
   
   Definition P_id_sqr (x8:Z) := 1.
   
   Definition P_id_cons (x8:Z) (x9:Z) := 0.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_0  := 0.
   
   Lemma P_id_active_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
   Proof.
     intros x9 x8.
     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_add_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_recip_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
   Proof.
     intros x9 x8.
     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_mark_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
   Proof.
     intros x9 x8.
     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_first_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_s_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
   Proof.
     intros x9 x8.
     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_terms_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
   Proof.
     intros x9 x8.
     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_dbl_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
   Proof.
     intros x9 x8.
     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_sqr_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
   Proof.
     intros x9 x8.
     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_cons_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_add_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_first_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_cons_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_nil_bounded : 0 <= P_id_nil .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_0_bounded : 0 <= P_id_0 .
   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_active P_id_add P_id_recip P_id_mark P_id_first 
      P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil P_id_0.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                    P_id_active (measure x8)
                   | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                    P_id_add (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                    P_id_recip (measure x8)
                   | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                    P_id_mark (measure x8)
                   | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                    P_id_first (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                    P_id_s (measure x8)
                   | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                    P_id_terms (measure x8)
                   | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                    P_id_dbl (measure x8)
                   | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                    P_id_sqr (measure x8)
                   | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                    P_id_cons (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                   | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_SQR (x8:Z) := 0.
   
   Definition P_id_DBL (x8:Z) := 0.
   
   Definition P_id_ACTIVE (x8:Z) := 0.
   
   Definition P_id_S (x8:Z) := 0.
   
   Definition P_id_CONS (x8:Z) (x9:Z) := 0.
   
   Definition P_id_TERMS (x8:Z) := 1* x8.
   
   Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
   
   Definition P_id_MARK (x8:Z) := 0.
   
   Definition P_id_ADD (x8:Z) (x9:Z) := 0.
   
   Definition P_id_RECIP (x8:Z) := 0.
   
   Lemma P_id_SQR_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
   Proof.
     intros x9 x8.
     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_DBL_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
   Proof.
     intros x9 x8.
     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_ACTIVE_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
   Proof.
     intros x9 x8.
     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_S_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
   Proof.
     intros x9 x8.
     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_CONS_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_TERMS_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
   Proof.
     intros x9 x8.
     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_FIRST_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_MARK_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
   Proof.
     intros x9 x8.
     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_ADD_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_RECIP_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
   Proof.
     intros x9 x8.
     intros [H_1 H_0].
     
     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_active P_id_add P_id_recip P_id_mark 
      P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
      P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
      P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                           P_id_SQR (measure x8)
                          | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                           P_id_DBL (measure x8)
                          | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                           P_id_ACTIVE (measure x8)
                          | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                           P_id_S (measure x8)
                          | (algebra.Alg.Term algebra.F.id_cons (x9::
                             x8::nil)) =>
                           P_id_CONS (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                           P_id_TERMS (measure x8)
                          | (algebra.Alg.Term algebra.F.id_first (x9::
                             x8::nil)) =>
                           P_id_FIRST (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                           P_id_MARK (measure x8)
                          | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                           P_id_ADD (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                           P_id_RECIP (measure x8)
                          | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_SQR_monotonic;assumption.
     intros ;apply P_id_DBL_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_S_monotonic;assumption.
     intros ;apply P_id_CONS_monotonic;assumption.
     intros ;apply P_id_TERMS_monotonic;assumption.
     intros ;apply P_id_FIRST_monotonic;assumption.
     intros ;apply P_id_MARK_monotonic;assumption.
     intros ;apply P_id_ADD_monotonic;assumption.
     intros ;apply P_id_RECIP_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_8_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_8_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_8_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_8_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_8_large  := WF_DP_R_xml_0_scc_8_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_8.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_8_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_8_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_8_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_8_large.
   Qed.
  End WF_DP_R_xml_0_scc_8.
  
  Definition wf_DP_R_xml_0_scc_8  := WF_DP_R_xml_0_scc_8.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_8 :
   forall x y, (DP_R_xml_0_scc_8 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_8).
    intros x' _ Hrec y h.
    
    inversion h;clear h;subst;
     constructor;intros _y _h;inversion _h;clear _h;subst;
      (eapply Hrec;econstructor eassumption)||
      ((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_8.
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_9  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <mark(first(X1_,X2_)),first(mark(X1_),mark(X2_))> *)
    | DP_R_xml_0_non_scc_9_0 :
     forall x8 x6 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
       DP_R_xml_0_non_scc_9 (algebra.Alg.Term algebra.F.id_first 
                             ((algebra.Alg.Term algebra.F.id_mark 
                             (x5::nil))::(algebra.Alg.Term algebra.F.id_mark 
                             (x6::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_9 :
   forall x y, 
    (DP_R_xml_0_non_scc_9 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;
      (eapply acc_DP_R_xml_0_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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_10  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <mark(dbl(X_)),dbl(mark(X_))> *)
    | DP_R_xml_0_non_scc_10_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) x8) ->
       DP_R_xml_0_non_scc_10 (algebra.Alg.Term algebra.F.id_dbl 
                              ((algebra.Alg.Term algebra.F.id_mark 
                              (x2::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_10 :
   forall x y, 
    (DP_R_xml_0_non_scc_10 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;
      (eapply acc_DP_R_xml_0_scc_2;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_11  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <mark(add(X1_,X2_)),add(mark(X1_),mark(X2_))> *)
    | DP_R_xml_0_non_scc_11_0 :
     forall x8 x6 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
       DP_R_xml_0_non_scc_11 (algebra.Alg.Term algebra.F.id_add 
                              ((algebra.Alg.Term algebra.F.id_mark 
                              (x5::nil))::(algebra.Alg.Term 
                              algebra.F.id_mark (x6::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_11 :
   forall x y, 
    (DP_R_xml_0_non_scc_11 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;
      (eapply acc_DP_R_xml_0_scc_3;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_12  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <mark(s(X_)),s(X_)> *)
    | DP_R_xml_0_non_scc_12_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 (algebra.Alg.Term algebra.F.id_s (x2::nil)) 
        x8) ->
       DP_R_xml_0_non_scc_12 (algebra.Alg.Term algebra.F.id_s (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_12 :
   forall x y, 
    (DP_R_xml_0_non_scc_12 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;
      (eapply acc_DP_R_xml_0_scc_4;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_13  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <mark(sqr(X_)),sqr(mark(X_))> *)
    | DP_R_xml_0_non_scc_13_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) x8) ->
       DP_R_xml_0_non_scc_13 (algebra.Alg.Term algebra.F.id_sqr 
                              ((algebra.Alg.Term algebra.F.id_mark 
                              (x2::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_13 :
   forall x y, 
    (DP_R_xml_0_non_scc_13 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;
      (eapply acc_DP_R_xml_0_scc_5;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_14  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <mark(recip(X_)),recip(mark(X_))> *)
    | DP_R_xml_0_non_scc_14_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_recip (x2::nil)) x8) ->
       DP_R_xml_0_non_scc_14 (algebra.Alg.Term algebra.F.id_recip 
                              ((algebra.Alg.Term algebra.F.id_mark 
                              (x2::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_14 :
   forall x y, 
    (DP_R_xml_0_non_scc_14 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;
      (eapply acc_DP_R_xml_0_scc_6;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_15  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <mark(cons(X1_,X2_)),cons(mark(X1_),X2_)> *)
    | DP_R_xml_0_non_scc_15_0 :
     forall x8 x6 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x8) ->
       DP_R_xml_0_non_scc_15 (algebra.Alg.Term algebra.F.id_cons 
                              ((algebra.Alg.Term algebra.F.id_mark 
                              (x5::nil))::x6::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_15 :
   forall x y, 
    (DP_R_xml_0_non_scc_15 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;
      (eapply acc_DP_R_xml_0_scc_7;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_16  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <mark(terms(X_)),terms(mark(X_))> *)
    | DP_R_xml_0_non_scc_16_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_terms (x2::nil)) x8) ->
       DP_R_xml_0_non_scc_16 (algebra.Alg.Term algebra.F.id_terms 
                              ((algebra.Alg.Term algebra.F.id_mark 
                              (x2::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_16 :
   forall x y, 
    (DP_R_xml_0_non_scc_16 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;
      (eapply acc_DP_R_xml_0_scc_8;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_17  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(first(s(X_),cons(Y_,Z_))),first(X_,Z_)> *)
    | DP_R_xml_0_non_scc_17_0 :
     forall x8 x4 x2 x3, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Term 
         algebra.F.id_s (x2::nil))::(algebra.Alg.Term algebra.F.id_cons (x3::
         x4::nil))::nil)) x8) ->
       DP_R_xml_0_non_scc_17 (algebra.Alg.Term algebra.F.id_first (x2::
                              x4::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_17 :
   forall x y, 
    (DP_R_xml_0_non_scc_17 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;
      (eapply acc_DP_R_xml_0_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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_18  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(first(s(X_),cons(Y_,Z_))),cons(Y_,first(X_,Z_))> *)
    | DP_R_xml_0_non_scc_18_0 :
     forall x8 x4 x2 x3, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Term 
         algebra.F.id_s (x2::nil))::(algebra.Alg.Term algebra.F.id_cons (x3::
         x4::nil))::nil)) x8) ->
       DP_R_xml_0_non_scc_18 (algebra.Alg.Term algebra.F.id_cons (x3::
                              (algebra.Alg.Term algebra.F.id_first (x2::
                              x4::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_18 :
   forall x y, 
    (DP_R_xml_0_non_scc_18 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;
      (eapply acc_DP_R_xml_0_scc_7;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_19  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(add(s(X_),Y_)),add(X_,Y_)> *)
    | DP_R_xml_0_non_scc_19_0 :
     forall x8 x2 x3, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term algebra.F.id_s 
         (x2::nil))::x3::nil)) x8) ->
       DP_R_xml_0_non_scc_19 (algebra.Alg.Term algebra.F.id_add (x2::
                              x3::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_19 :
   forall x y, 
    (DP_R_xml_0_non_scc_19 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;
      (eapply acc_DP_R_xml_0_scc_3;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_20  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(add(s(X_),Y_)),s(add(X_,Y_))> *)
    | DP_R_xml_0_non_scc_20_0 :
     forall x8 x2 x3, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term algebra.F.id_s 
         (x2::nil))::x3::nil)) x8) ->
       DP_R_xml_0_non_scc_20 (algebra.Alg.Term algebra.F.id_s 
                              ((algebra.Alg.Term algebra.F.id_add (x2::
                              x3::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_20 :
   forall x y, 
    (DP_R_xml_0_non_scc_20 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;
      (eapply acc_DP_R_xml_0_scc_4;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_21  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(dbl(s(X_))),dbl(X_)> *)
    | DP_R_xml_0_non_scc_21_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_s 
         (x2::nil))::nil)) x8) ->
       DP_R_xml_0_non_scc_21 (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_21 :
   forall x y, 
    (DP_R_xml_0_non_scc_21 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;
      (eapply acc_DP_R_xml_0_scc_2;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_22  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(dbl(s(X_))),s(dbl(X_))> *)
    | DP_R_xml_0_non_scc_22_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_s 
         (x2::nil))::nil)) x8) ->
       DP_R_xml_0_non_scc_22 (algebra.Alg.Term algebra.F.id_s 
                              ((algebra.Alg.Term algebra.F.id_dbl 
                              (x2::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_22 :
   forall x y, 
    (DP_R_xml_0_non_scc_22 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;
      (eapply acc_DP_R_xml_0_scc_4;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_23  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(dbl(s(X_))),s(s(dbl(X_)))> *)
    | DP_R_xml_0_non_scc_23_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_s 
         (x2::nil))::nil)) x8) ->
       DP_R_xml_0_non_scc_23 (algebra.Alg.Term algebra.F.id_s 
                              ((algebra.Alg.Term algebra.F.id_s 
                              ((algebra.Alg.Term algebra.F.id_dbl 
                              (x2::nil))::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_23 :
   forall x y, 
    (DP_R_xml_0_non_scc_23 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;
      (eapply acc_DP_R_xml_0_scc_4;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_24  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(sqr(s(X_))),dbl(X_)> *)
    | DP_R_xml_0_non_scc_24_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_s 
         (x2::nil))::nil)) x8) ->
       DP_R_xml_0_non_scc_24 (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_24 :
   forall x y, 
    (DP_R_xml_0_non_scc_24 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;
      (eapply acc_DP_R_xml_0_scc_2;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_25  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(sqr(s(X_))),sqr(X_)> *)
    | DP_R_xml_0_non_scc_25_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_s 
         (x2::nil))::nil)) x8) ->
       DP_R_xml_0_non_scc_25 (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_25 :
   forall x y, 
    (DP_R_xml_0_non_scc_25 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;
      (eapply acc_DP_R_xml_0_scc_5;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_26  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(sqr(s(X_))),add(sqr(X_),dbl(X_))> *)
    | DP_R_xml_0_non_scc_26_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_s 
         (x2::nil))::nil)) x8) ->
       DP_R_xml_0_non_scc_26 (algebra.Alg.Term algebra.F.id_add 
                              ((algebra.Alg.Term algebra.F.id_sqr 
                              (x2::nil))::(algebra.Alg.Term algebra.F.id_dbl 
                              (x2::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_26 :
   forall x y, 
    (DP_R_xml_0_non_scc_26 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;
      (eapply acc_DP_R_xml_0_scc_3;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_27  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(sqr(s(X_))),s(add(sqr(X_),dbl(X_)))> *)
    | DP_R_xml_0_non_scc_27_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_s 
         (x2::nil))::nil)) x8) ->
       DP_R_xml_0_non_scc_27 (algebra.Alg.Term algebra.F.id_s 
                              ((algebra.Alg.Term algebra.F.id_add 
                              ((algebra.Alg.Term algebra.F.id_sqr 
                              (x2::nil))::(algebra.Alg.Term algebra.F.id_dbl 
                              (x2::nil))::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_27 :
   forall x y, 
    (DP_R_xml_0_non_scc_27 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;
      (eapply acc_DP_R_xml_0_scc_4;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_28  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(terms(N_)),s(N_)> *)
    | DP_R_xml_0_non_scc_28_0 :
     forall x8 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
       DP_R_xml_0_non_scc_28 (algebra.Alg.Term algebra.F.id_s (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_28 :
   forall x y, 
    (DP_R_xml_0_non_scc_28 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;
      (eapply acc_DP_R_xml_0_scc_4;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_29  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(terms(N_)),terms(s(N_))> *)
    | DP_R_xml_0_non_scc_29_0 :
     forall x8 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
       DP_R_xml_0_non_scc_29 (algebra.Alg.Term algebra.F.id_terms 
                              ((algebra.Alg.Term algebra.F.id_s 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_29 :
   forall x y, 
    (DP_R_xml_0_non_scc_29 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;
      (eapply acc_DP_R_xml_0_scc_8;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_30  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(terms(N_)),sqr(N_)> *)
    | DP_R_xml_0_non_scc_30_0 :
     forall x8 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
       DP_R_xml_0_non_scc_30 (algebra.Alg.Term algebra.F.id_sqr (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_30 :
   forall x y, 
    (DP_R_xml_0_non_scc_30 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;
      (eapply acc_DP_R_xml_0_scc_5;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_31  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(terms(N_)),recip(sqr(N_))> *)
    | DP_R_xml_0_non_scc_31_0 :
     forall x8 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
       DP_R_xml_0_non_scc_31 (algebra.Alg.Term algebra.F.id_recip 
                              ((algebra.Alg.Term algebra.F.id_sqr 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_31 :
   forall x y, 
    (DP_R_xml_0_non_scc_31 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;
      (eapply acc_DP_R_xml_0_scc_6;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_32  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(terms(N_)),cons(recip(sqr(N_)),terms(s(N_)))> *)
    | DP_R_xml_0_non_scc_32_0 :
     forall x8 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
       DP_R_xml_0_non_scc_32 (algebra.Alg.Term algebra.F.id_cons 
                              ((algebra.Alg.Term algebra.F.id_recip 
                              ((algebra.Alg.Term algebra.F.id_sqr 
                              (x1::nil))::nil))::(algebra.Alg.Term 
                              algebra.F.id_terms ((algebra.Alg.Term 
                              algebra.F.id_s (x1::nil))::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_32 :
   forall x y, 
    (DP_R_xml_0_non_scc_32 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;
      (eapply acc_DP_R_xml_0_scc_7;
        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' ))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_33  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <mark(nil),active(nil)> *)
    | DP_R_xml_0_non_scc_33_0 :
     forall x8, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 (algebra.Alg.Term algebra.F.id_nil nil) 
        x8) ->
       DP_R_xml_0_non_scc_33 (algebra.Alg.Term algebra.F.id_active 
                              ((algebra.Alg.Term algebra.F.id_nil nil)::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_33 :
   forall x y, 
    (DP_R_xml_0_non_scc_33 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;
      (eapply acc_DP_R_xml_0_non_scc_32;
        econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
      ((eapply acc_DP_R_xml_0_non_scc_31;
         econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
       ((eapply acc_DP_R_xml_0_non_scc_30;
          econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
        ((eapply acc_DP_R_xml_0_non_scc_29;
           econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
         ((eapply acc_DP_R_xml_0_non_scc_28;
            econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
          ((eapply acc_DP_R_xml_0_non_scc_27;
             econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
           ((eapply acc_DP_R_xml_0_non_scc_26;
              econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
            ((eapply acc_DP_R_xml_0_non_scc_25;
               econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
             ((eapply acc_DP_R_xml_0_non_scc_24;
                econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
              ((eapply acc_DP_R_xml_0_non_scc_23;
                 econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
               ((eapply acc_DP_R_xml_0_non_scc_22;
                  econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                ((eapply acc_DP_R_xml_0_non_scc_21;
                   econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                 ((eapply acc_DP_R_xml_0_non_scc_20;
                    econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                  ((eapply acc_DP_R_xml_0_non_scc_19;
                     econstructor 
                     (eassumption)||(algebra.Alg_ext.star_refl' ))||
                   ((eapply acc_DP_R_xml_0_non_scc_18;
                      econstructor 
                      (eassumption)||(algebra.Alg_ext.star_refl' ))||
                    ((eapply acc_DP_R_xml_0_non_scc_17;
                       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' )))))))))))))))))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_34  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(first(0,X_)),mark(nil)> *)
    | DP_R_xml_0_non_scc_34_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Term 
         algebra.F.id_0 nil)::x2::nil)) x8) ->
       DP_R_xml_0_non_scc_34 (algebra.Alg.Term algebra.F.id_mark 
                              ((algebra.Alg.Term algebra.F.id_nil nil)::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_34 :
   forall x y, 
    (DP_R_xml_0_non_scc_34 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;
      (eapply acc_DP_R_xml_0_non_scc_33;
        econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
      ((eapply acc_DP_R_xml_0_non_scc_16;
         econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
       ((eapply acc_DP_R_xml_0_non_scc_15;
          econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
        ((eapply acc_DP_R_xml_0_non_scc_14;
           econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
         ((eapply acc_DP_R_xml_0_non_scc_13;
            econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
          ((eapply acc_DP_R_xml_0_non_scc_12;
             econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
           ((eapply acc_DP_R_xml_0_non_scc_11;
              econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
            ((eapply acc_DP_R_xml_0_non_scc_10;
               econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
             ((eapply acc_DP_R_xml_0_non_scc_9;
                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' ))))))))))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_35  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <mark(0),active(0)> *)
    | DP_R_xml_0_non_scc_35_0 :
     forall x8, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 (algebra.Alg.Term algebra.F.id_0 nil) 
        x8) ->
       DP_R_xml_0_non_scc_35 (algebra.Alg.Term algebra.F.id_active 
                              ((algebra.Alg.Term algebra.F.id_0 nil)::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_35 :
   forall x y, 
    (DP_R_xml_0_non_scc_35 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;
      (eapply acc_DP_R_xml_0_non_scc_32;
        econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
      ((eapply acc_DP_R_xml_0_non_scc_31;
         econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
       ((eapply acc_DP_R_xml_0_non_scc_30;
          econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
        ((eapply acc_DP_R_xml_0_non_scc_29;
           econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
         ((eapply acc_DP_R_xml_0_non_scc_28;
            econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
          ((eapply acc_DP_R_xml_0_non_scc_27;
             econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
           ((eapply acc_DP_R_xml_0_non_scc_26;
              econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
            ((eapply acc_DP_R_xml_0_non_scc_25;
               econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
             ((eapply acc_DP_R_xml_0_non_scc_24;
                econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
              ((eapply acc_DP_R_xml_0_non_scc_23;
                 econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
               ((eapply acc_DP_R_xml_0_non_scc_22;
                  econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                ((eapply acc_DP_R_xml_0_non_scc_21;
                   econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                 ((eapply acc_DP_R_xml_0_non_scc_20;
                    econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                  ((eapply acc_DP_R_xml_0_non_scc_19;
                     econstructor 
                     (eassumption)||(algebra.Alg_ext.star_refl' ))||
                   ((eapply acc_DP_R_xml_0_non_scc_18;
                      econstructor 
                      (eassumption)||(algebra.Alg_ext.star_refl' ))||
                    ((eapply acc_DP_R_xml_0_non_scc_17;
                       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' )))))))))))))))))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_36  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(dbl(0)),mark(0)> *)
    | DP_R_xml_0_non_scc_36_0 :
     forall x8, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_0 
         nil)::nil)) x8) ->
       DP_R_xml_0_non_scc_36 (algebra.Alg.Term algebra.F.id_mark 
                              ((algebra.Alg.Term algebra.F.id_0 nil)::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_36 :
   forall x y, 
    (DP_R_xml_0_non_scc_36 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;
      (eapply acc_DP_R_xml_0_non_scc_35;
        econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
      ((eapply acc_DP_R_xml_0_non_scc_16;
         econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
       ((eapply acc_DP_R_xml_0_non_scc_15;
          econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
        ((eapply acc_DP_R_xml_0_non_scc_14;
           econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
         ((eapply acc_DP_R_xml_0_non_scc_13;
            econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
          ((eapply acc_DP_R_xml_0_non_scc_12;
             econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
           ((eapply acc_DP_R_xml_0_non_scc_11;
              econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
            ((eapply acc_DP_R_xml_0_non_scc_10;
               econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
             ((eapply acc_DP_R_xml_0_non_scc_9;
                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' ))))))))))).
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_37  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(sqr(0)),mark(0)> *)
    | DP_R_xml_0_non_scc_37_0 :
     forall x8, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_0 
         nil)::nil)) x8) ->
       DP_R_xml_0_non_scc_37 (algebra.Alg.Term algebra.F.id_mark 
                              ((algebra.Alg.Term algebra.F.id_0 nil)::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_37 :
   forall x y, 
    (DP_R_xml_0_non_scc_37 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;
      (eapply acc_DP_R_xml_0_non_scc_35;
        econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
      ((eapply acc_DP_R_xml_0_non_scc_16;
         econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
       ((eapply acc_DP_R_xml_0_non_scc_15;
          econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
        ((eapply acc_DP_R_xml_0_non_scc_14;
           econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
         ((eapply acc_DP_R_xml_0_non_scc_13;
            econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
          ((eapply acc_DP_R_xml_0_non_scc_12;
             econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
           ((eapply acc_DP_R_xml_0_non_scc_11;
              econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
            ((eapply acc_DP_R_xml_0_non_scc_10;
               econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
             ((eapply acc_DP_R_xml_0_non_scc_9;
                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' ))))))))))).
  Qed.
  
  
  Inductive DP_R_xml_0_scc_38  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <mark(terms(X_)),active(terms(mark(X_)))> *)
    | DP_R_xml_0_scc_38_0 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_terms (x2::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_active 
                          ((algebra.Alg.Term algebra.F.id_terms 
                          ((algebra.Alg.Term algebra.F.id_mark 
                          (x2::nil))::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    
     (* <active(terms(N_)),mark(cons(recip(sqr(N_)),terms(s(N_))))> *)
    | DP_R_xml_0_scc_38_1 :
     forall x8 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_mark 
                          ((algebra.Alg.Term algebra.F.id_cons 
                          ((algebra.Alg.Term algebra.F.id_recip 
                          ((algebra.Alg.Term algebra.F.id_sqr 
                          (x1::nil))::nil))::(algebra.Alg.Term 
                          algebra.F.id_terms ((algebra.Alg.Term 
                          algebra.F.id_s (x1::nil))::nil))::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
     (* <mark(terms(X_)),mark(X_)> *)
    | DP_R_xml_0_scc_38_2 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_terms (x2::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_mark (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    
     (* <mark(cons(X1_,X2_)),active(cons(mark(X1_),X2_))> *)
    | DP_R_xml_0_scc_38_3 :
     forall x8 x6 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_active 
                          ((algebra.Alg.Term algebra.F.id_cons 
                          ((algebra.Alg.Term algebra.F.id_mark (x5::nil))::
                          x6::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    
     (* <active(sqr(s(X_))),mark(s(add(sqr(X_),dbl(X_))))> *)
    | DP_R_xml_0_scc_38_4 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term algebra.F.id_s 
         (x2::nil))::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_mark 
                          ((algebra.Alg.Term algebra.F.id_s 
                          ((algebra.Alg.Term algebra.F.id_add 
                          ((algebra.Alg.Term algebra.F.id_sqr (x2::nil))::
                          (algebra.Alg.Term algebra.F.id_dbl 
                          (x2::nil))::nil))::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
     (* <mark(cons(X1_,X2_)),mark(X1_)> *)
    | DP_R_xml_0_scc_38_5 :
     forall x8 x6 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     (* <mark(recip(X_)),active(recip(mark(X_)))> *)
    | DP_R_xml_0_scc_38_6 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_recip (x2::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_active 
                          ((algebra.Alg.Term algebra.F.id_recip 
                          ((algebra.Alg.Term algebra.F.id_mark 
                          (x2::nil))::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     (* <active(dbl(s(X_))),mark(s(s(dbl(X_))))> *)
    | DP_R_xml_0_scc_38_7 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term algebra.F.id_s 
         (x2::nil))::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_mark 
                          ((algebra.Alg.Term algebra.F.id_s 
                          ((algebra.Alg.Term algebra.F.id_s 
                          ((algebra.Alg.Term algebra.F.id_dbl 
                          (x2::nil))::nil))::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
     (* <mark(recip(X_)),mark(X_)> *)
    | DP_R_xml_0_scc_38_8 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_recip (x2::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_mark (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     (* <mark(sqr(X_)),active(sqr(mark(X_)))> *)
    | DP_R_xml_0_scc_38_9 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_active 
                          ((algebra.Alg.Term algebra.F.id_sqr 
                          ((algebra.Alg.Term algebra.F.id_mark 
                          (x2::nil))::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     (* <active(add(0,X_)),mark(X_)> *)
    | DP_R_xml_0_scc_38_10 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term algebra.F.id_0 
         nil)::x2::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_mark (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
     (* <mark(sqr(X_)),mark(X_)> *)
    | DP_R_xml_0_scc_38_11 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_mark (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     (* <mark(s(X_)),active(s(X_))> *)
    | DP_R_xml_0_scc_38_12 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 (algebra.Alg.Term algebra.F.id_s (x2::nil)) 
        x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_active 
                          ((algebra.Alg.Term algebra.F.id_s (x2::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    
     (* <active(add(s(X_),Y_)),mark(s(add(X_,Y_)))> *)
    | DP_R_xml_0_scc_38_13 :
     forall x8 x2 x3, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term algebra.F.id_s 
         (x2::nil))::x3::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_mark 
                          ((algebra.Alg.Term algebra.F.id_s 
                          ((algebra.Alg.Term algebra.F.id_add (x2::
                          x3::nil))::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
    
     (* <mark(add(X1_,X2_)),active(add(mark(X1_),mark(X2_)))> *)
    | DP_R_xml_0_scc_38_14 :
     forall x8 x6 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_active 
                          ((algebra.Alg.Term algebra.F.id_add 
                          ((algebra.Alg.Term algebra.F.id_mark (x5::nil))::
                          (algebra.Alg.Term algebra.F.id_mark 
                          (x6::nil))::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    
     (* <active(first(s(X_),cons(Y_,Z_))),mark(cons(Y_,first(X_,Z_)))> *)
    | DP_R_xml_0_scc_38_15 :
     forall x8 x4 x2 x3, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Term 
         algebra.F.id_s (x2::nil))::(algebra.Alg.Term algebra.F.id_cons (x3::
         x4::nil))::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_mark 
                          ((algebra.Alg.Term algebra.F.id_cons (x3::
                          (algebra.Alg.Term algebra.F.id_first (x2::
                          x4::nil))::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x8::nil))
     (* <mark(add(X1_,X2_)),mark(X1_)> *)
    | DP_R_xml_0_scc_38_16 :
     forall x8 x6 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     (* <mark(add(X1_,X2_)),mark(X2_)> *)
    | DP_R_xml_0_scc_38_17 :
     forall x8 x6 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_mark (x6::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     (* <mark(dbl(X_)),active(dbl(mark(X_)))> *)
    | DP_R_xml_0_scc_38_18 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_active 
                          ((algebra.Alg.Term algebra.F.id_dbl 
                          ((algebra.Alg.Term algebra.F.id_mark 
                          (x2::nil))::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     (* <mark(dbl(X_)),mark(X_)> *)
    | DP_R_xml_0_scc_38_19 :
     forall x8 x2, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_mark (x2::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    
     (* <mark(first(X1_,X2_)),active(first(mark(X1_),mark(X2_)))> *)
    | DP_R_xml_0_scc_38_20 :
     forall x8 x6 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_active 
                          ((algebra.Alg.Term algebra.F.id_first 
                          ((algebra.Alg.Term algebra.F.id_mark (x5::nil))::
                          (algebra.Alg.Term algebra.F.id_mark 
                          (x6::nil))::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     (* <mark(first(X1_,X2_)),mark(X1_)> *)
    | DP_R_xml_0_scc_38_21 :
     forall x8 x6 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     (* <mark(first(X1_,X2_)),mark(X2_)> *)
    | DP_R_xml_0_scc_38_22 :
     forall x8 x6 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
       DP_R_xml_0_scc_38 (algebra.Alg.Term algebra.F.id_mark (x6::nil)) 
        (algebra.Alg.Term algebra.F.id_mark (x8::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_38.
   Inductive DP_R_xml_0_scc_38_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <mark(terms(X_)),active(terms(mark(X_)))> *)
     | DP_R_xml_0_scc_38_large_0 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_terms (x2::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_active 
                                 ((algebra.Alg.Term algebra.F.id_terms 
                                 ((algebra.Alg.Term algebra.F.id_mark 
                                 (x2::nil))::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     
      (* <active(terms(N_)),mark(cons(recip(sqr(N_)),terms(s(N_))))> *)
     | DP_R_xml_0_scc_38_large_1 :
      forall x8 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_mark 
                                 ((algebra.Alg.Term algebra.F.id_cons 
                                 ((algebra.Alg.Term algebra.F.id_recip 
                                 ((algebra.Alg.Term algebra.F.id_sqr 
                                 (x1::nil))::nil))::(algebra.Alg.Term 
                                 algebra.F.id_terms ((algebra.Alg.Term 
                                 algebra.F.id_s 
                                 (x1::nil))::nil))::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_active (x8::nil))
      (* <mark(terms(X_)),mark(X_)> *)
     | DP_R_xml_0_scc_38_large_2 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_terms (x2::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_mark 
                                 (x2::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     
      (* <mark(cons(X1_,X2_)),active(cons(mark(X1_),X2_))> *)
     | DP_R_xml_0_scc_38_large_3 :
      forall x8 x6 x5, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_active 
                                 ((algebra.Alg.Term algebra.F.id_cons 
                                 ((algebra.Alg.Term algebra.F.id_mark 
                                 (x5::nil))::x6::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     
      (* <active(sqr(s(X_))),mark(s(add(sqr(X_),dbl(X_))))> *)
     | DP_R_xml_0_scc_38_large_4 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term 
          algebra.F.id_s (x2::nil))::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_mark 
                                 ((algebra.Alg.Term algebra.F.id_s 
                                 ((algebra.Alg.Term algebra.F.id_add 
                                 ((algebra.Alg.Term algebra.F.id_sqr 
                                 (x2::nil))::(algebra.Alg.Term 
                                 algebra.F.id_dbl 
                                 (x2::nil))::nil))::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_active (x8::nil))
      (* <mark(cons(X1_,X2_)),mark(X1_)> *)
     | DP_R_xml_0_scc_38_large_5 :
      forall x8 x6 x5, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_mark 
                                 (x5::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     
      (* <mark(recip(X_)),active(recip(mark(X_)))> *)
     | DP_R_xml_0_scc_38_large_6 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_recip (x2::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_active 
                                 ((algebra.Alg.Term algebra.F.id_recip 
                                 ((algebra.Alg.Term algebra.F.id_mark 
                                 (x2::nil))::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     
      (* <active(dbl(s(X_))),mark(s(s(dbl(X_))))> *)
     | DP_R_xml_0_scc_38_large_7 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term 
          algebra.F.id_s (x2::nil))::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_mark 
                                 ((algebra.Alg.Term algebra.F.id_s 
                                 ((algebra.Alg.Term algebra.F.id_s 
                                 ((algebra.Alg.Term algebra.F.id_dbl 
                                 (x2::nil))::nil))::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_active (x8::nil))
      (* <mark(recip(X_)),mark(X_)> *)
     | DP_R_xml_0_scc_38_large_8 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_recip (x2::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_mark 
                                 (x2::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     
      (* <mark(sqr(X_)),active(sqr(mark(X_)))> *)
     | DP_R_xml_0_scc_38_large_9 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_active 
                                 ((algebra.Alg.Term algebra.F.id_sqr 
                                 ((algebra.Alg.Term algebra.F.id_mark 
                                 (x2::nil))::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      (* <active(add(0,X_)),mark(X_)> *)
     | DP_R_xml_0_scc_38_large_10 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term 
          algebra.F.id_0 nil)::x2::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_mark 
                                 (x2::nil)) 
         (algebra.Alg.Term algebra.F.id_active (x8::nil))
      (* <mark(sqr(X_)),mark(X_)> *)
     | DP_R_xml_0_scc_38_large_11 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_mark 
                                 (x2::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     
      (* <active(add(s(X_),Y_)),mark(s(add(X_,Y_)))> *)
     | DP_R_xml_0_scc_38_large_12 :
      forall x8 x2 x3, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term 
          algebra.F.id_s (x2::nil))::x3::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_mark 
                                 ((algebra.Alg.Term algebra.F.id_s 
                                 ((algebra.Alg.Term algebra.F.id_add (x2::
                                 x3::nil))::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_active (x8::nil))
     
      (* <mark(add(X1_,X2_)),active(add(mark(X1_),mark(X2_)))> *)
     | DP_R_xml_0_scc_38_large_13 :
      forall x8 x6 x5, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_active 
                                 ((algebra.Alg.Term algebra.F.id_add 
                                 ((algebra.Alg.Term algebra.F.id_mark 
                                 (x5::nil))::(algebra.Alg.Term 
                                 algebra.F.id_mark (x6::nil))::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     
      (* <active(first(s(X_),cons(Y_,Z_))),mark(cons(Y_,first(X_,Z_)))> *)
     | DP_R_xml_0_scc_38_large_14 :
      forall x8 x4 x2 x3, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Term 
          algebra.F.id_s (x2::nil))::(algebra.Alg.Term algebra.F.id_cons 
          (x3::x4::nil))::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_mark 
                                 ((algebra.Alg.Term algebra.F.id_cons (x3::
                                 (algebra.Alg.Term algebra.F.id_first (x2::
                                 x4::nil))::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_active (x8::nil))
      (* <mark(add(X1_,X2_)),mark(X1_)> *)
     | DP_R_xml_0_scc_38_large_15 :
      forall x8 x6 x5, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_mark 
                                 (x5::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      (* <mark(add(X1_,X2_)),mark(X2_)> *)
     | DP_R_xml_0_scc_38_large_16 :
      forall x8 x6 x5, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_mark 
                                 (x6::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     
      (* <mark(dbl(X_)),active(dbl(mark(X_)))> *)
     | DP_R_xml_0_scc_38_large_17 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_active 
                                 ((algebra.Alg.Term algebra.F.id_dbl 
                                 ((algebra.Alg.Term algebra.F.id_mark 
                                 (x2::nil))::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      (* <mark(dbl(X_)),mark(X_)> *)
     | DP_R_xml_0_scc_38_large_18 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_mark 
                                 (x2::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     
      (* <mark(first(X1_,X2_)),active(first(mark(X1_),mark(X2_)))> *)
     | DP_R_xml_0_scc_38_large_19 :
      forall x8 x6 x5, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_active 
                                 ((algebra.Alg.Term algebra.F.id_first 
                                 ((algebra.Alg.Term algebra.F.id_mark 
                                 (x5::nil))::(algebra.Alg.Term 
                                 algebra.F.id_mark (x6::nil))::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      (* <mark(first(X1_,X2_)),mark(X1_)> *)
     | DP_R_xml_0_scc_38_large_20 :
      forall x8 x6 x5, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_mark 
                                 (x5::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      (* <mark(first(X1_,X2_)),mark(X2_)> *)
     | DP_R_xml_0_scc_38_large_21 :
      forall x8 x6 x5, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
        DP_R_xml_0_scc_38_large (algebra.Alg.Term algebra.F.id_mark 
                                 (x6::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_38_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <mark(s(X_)),active(s(X_))> *)
     | DP_R_xml_0_scc_38_strict_0 :
      forall x8 x2, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_s (x2::nil)) x8) ->
        DP_R_xml_0_scc_38_strict (algebra.Alg.Term algebra.F.id_active 
                                  ((algebra.Alg.Term algebra.F.id_s 
                                  (x2::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_mark (x8::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_38_large.
    Inductive DP_R_xml_0_scc_38_large_large  :
     algebra.Alg.term ->algebra.Alg.term ->Prop := 
       (* <mark(terms(X_)),active(terms(mark(X_)))> *)
      | DP_R_xml_0_scc_38_large_large_0 :
       forall x8 x2, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_terms (x2::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_active 
                                        ((algebra.Alg.Term 
                                        algebra.F.id_terms 
                                        ((algebra.Alg.Term algebra.F.id_mark 
                                        (x2::nil))::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      
       (* <active(terms(N_)),mark(cons(recip(sqr(N_)),terms(s(N_))))> *)
      | DP_R_xml_0_scc_38_large_large_1 :
       forall x8 x1, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_mark 
                                        ((algebra.Alg.Term algebra.F.id_cons 
                                        ((algebra.Alg.Term 
                                        algebra.F.id_recip 
                                        ((algebra.Alg.Term algebra.F.id_sqr 
                                        (x1::nil))::nil))::(algebra.Alg.Term 
                                        algebra.F.id_terms 
                                        ((algebra.Alg.Term algebra.F.id_s 
                                        (x1::nil))::nil))::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_active (x8::nil))
       (* <mark(terms(X_)),mark(X_)> *)
      | DP_R_xml_0_scc_38_large_large_2 :
       forall x8 x2, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_terms (x2::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_mark 
                                        (x2::nil)) 
          (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      
       (* <active(sqr(s(X_))),mark(s(add(sqr(X_),dbl(X_))))> *)
      | DP_R_xml_0_scc_38_large_large_3 :
       forall x8 x2, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term 
           algebra.F.id_s (x2::nil))::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_mark 
                                        ((algebra.Alg.Term algebra.F.id_s 
                                        ((algebra.Alg.Term algebra.F.id_add 
                                        ((algebra.Alg.Term algebra.F.id_sqr 
                                        (x2::nil))::(algebra.Alg.Term 
                                        algebra.F.id_dbl 
                                        (x2::nil))::nil))::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_active (x8::nil))
      
       (* <mark(cons(X1_,X2_)),mark(X1_)> *)
      | DP_R_xml_0_scc_38_large_large_4 :
       forall x8 x6 x5, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_mark 
                                        (x5::nil)) 
          (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      
       (* <active(dbl(s(X_))),mark(s(s(dbl(X_))))> *)
      | DP_R_xml_0_scc_38_large_large_5 :
       forall x8 x2, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term 
           algebra.F.id_s (x2::nil))::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_mark 
                                        ((algebra.Alg.Term algebra.F.id_s 
                                        ((algebra.Alg.Term algebra.F.id_s 
                                        ((algebra.Alg.Term algebra.F.id_dbl 
                                        (x2::nil))::nil))::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_active (x8::nil))
       (* <mark(recip(X_)),mark(X_)> *)
      | DP_R_xml_0_scc_38_large_large_6 :
       forall x8 x2, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_recip (x2::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_mark 
                                        (x2::nil)) 
          (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      
       (* <mark(sqr(X_)),active(sqr(mark(X_)))> *)
      | DP_R_xml_0_scc_38_large_large_7 :
       forall x8 x2, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_active 
                                        ((algebra.Alg.Term algebra.F.id_sqr 
                                        ((algebra.Alg.Term algebra.F.id_mark 
                                        (x2::nil))::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_mark (x8::nil))
       (* <active(add(0,X_)),mark(X_)> *)
      | DP_R_xml_0_scc_38_large_large_8 :
       forall x8 x2, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term 
           algebra.F.id_0 nil)::x2::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_mark 
                                        (x2::nil)) 
          (algebra.Alg.Term algebra.F.id_active (x8::nil))
       (* <mark(sqr(X_)),mark(X_)> *)
      | DP_R_xml_0_scc_38_large_large_9 :
       forall x8 x2, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_mark 
                                        (x2::nil)) 
          (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      
       (* <active(add(s(X_),Y_)),mark(s(add(X_,Y_)))> *)
      | DP_R_xml_0_scc_38_large_large_10 :
       forall x8 x2 x3, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term 
           algebra.F.id_s (x2::nil))::x3::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_mark 
                                        ((algebra.Alg.Term algebra.F.id_s 
                                        ((algebra.Alg.Term algebra.F.id_add 
                                        (x2::x3::nil))::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_active (x8::nil))
      
       (* <mark(add(X1_,X2_)),active(add(mark(X1_),mark(X2_)))> *)
      | DP_R_xml_0_scc_38_large_large_11 :
       forall x8 x6 x5, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_active 
                                        ((algebra.Alg.Term algebra.F.id_add 
                                        ((algebra.Alg.Term algebra.F.id_mark 
                                        (x5::nil))::(algebra.Alg.Term 
                                        algebra.F.id_mark 
                                        (x6::nil))::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      
       (* <active(first(s(X_),cons(Y_,Z_))),mark(cons(Y_,first(X_,Z_)))> *)
      | DP_R_xml_0_scc_38_large_large_12 :
       forall x8 x4 x2 x3, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Term 
           algebra.F.id_s (x2::nil))::(algebra.Alg.Term algebra.F.id_cons 
           (x3::x4::nil))::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_mark 
                                        ((algebra.Alg.Term algebra.F.id_cons 
                                        (x3::(algebra.Alg.Term 
                                        algebra.F.id_first (x2::
                                        x4::nil))::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_active (x8::nil))
      
       (* <mark(add(X1_,X2_)),mark(X1_)> *)
      | DP_R_xml_0_scc_38_large_large_13 :
       forall x8 x6 x5, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_mark 
                                        (x5::nil)) 
          (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      
       (* <mark(add(X1_,X2_)),mark(X2_)> *)
      | DP_R_xml_0_scc_38_large_large_14 :
       forall x8 x6 x5, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_mark 
                                        (x6::nil)) 
          (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      
       (* <mark(dbl(X_)),active(dbl(mark(X_)))> *)
      | DP_R_xml_0_scc_38_large_large_15 :
       forall x8 x2, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_active 
                                        ((algebra.Alg.Term algebra.F.id_dbl 
                                        ((algebra.Alg.Term algebra.F.id_mark 
                                        (x2::nil))::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_mark (x8::nil))
       (* <mark(dbl(X_)),mark(X_)> *)
      | DP_R_xml_0_scc_38_large_large_16 :
       forall x8 x2, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_mark 
                                        (x2::nil)) 
          (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      
       (* <mark(first(X1_,X2_)),active(first(mark(X1_),mark(X2_)))> *)
      | DP_R_xml_0_scc_38_large_large_17 :
       forall x8 x6 x5, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_active 
                                        ((algebra.Alg.Term 
                                        algebra.F.id_first 
                                        ((algebra.Alg.Term algebra.F.id_mark 
                                        (x5::nil))::(algebra.Alg.Term 
                                        algebra.F.id_mark 
                                        (x6::nil))::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      
       (* <mark(first(X1_,X2_)),mark(X1_)> *)
      | DP_R_xml_0_scc_38_large_large_18 :
       forall x8 x6 x5, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_mark 
                                        (x5::nil)) 
          (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      
       (* <mark(first(X1_,X2_)),mark(X2_)> *)
      | DP_R_xml_0_scc_38_large_large_19 :
       forall x8 x6 x5, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
         DP_R_xml_0_scc_38_large_large (algebra.Alg.Term algebra.F.id_mark 
                                        (x6::nil)) 
          (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    .
    
    
    Inductive DP_R_xml_0_scc_38_large_strict  :
     algebra.Alg.term ->algebra.Alg.term ->Prop := 
       (* <mark(cons(X1_,X2_)),active(cons(mark(X1_),X2_))> *)
      | DP_R_xml_0_scc_38_large_strict_0 :
       forall x8 x6 x5, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x8) ->
         DP_R_xml_0_scc_38_large_strict (algebra.Alg.Term 
                                         algebra.F.id_active 
                                         ((algebra.Alg.Term 
                                         algebra.F.id_cons 
                                         ((algebra.Alg.Term 
                                         algebra.F.id_mark (x5::nil))::
                                         x6::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      
       (* <mark(recip(X_)),active(recip(mark(X_)))> *)
      | DP_R_xml_0_scc_38_large_strict_1 :
       forall x8 x2, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_recip (x2::nil)) x8) ->
         DP_R_xml_0_scc_38_large_strict (algebra.Alg.Term 
                                         algebra.F.id_active 
                                         ((algebra.Alg.Term 
                                         algebra.F.id_recip 
                                         ((algebra.Alg.Term 
                                         algebra.F.id_mark 
                                         (x2::nil))::nil))::nil)) 
          (algebra.Alg.Term algebra.F.id_mark (x8::nil))
    .
    
    
    Module WF_DP_R_xml_0_scc_38_large_large.
     Inductive DP_R_xml_0_scc_38_large_large_large  :
      algebra.Alg.term ->algebra.Alg.term ->Prop := 
        (* <mark(terms(X_)),active(terms(mark(X_)))> *)
       | DP_R_xml_0_scc_38_large_large_large_0 :
        forall x8 x2, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_terms (x2::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_large (algebra.Alg.Term 
                                               algebra.F.id_active 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_terms 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_mark 
                                               (x2::nil))::nil))::nil)) 
           (algebra.Alg.Term algebra.F.id_mark (x8::nil))
       
        (* <active(terms(N_)),mark(cons(recip(sqr(N_)),terms(s(N_))))> *)
       | DP_R_xml_0_scc_38_large_large_large_1 :
        forall x8 x1, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_large (algebra.Alg.Term 
                                               algebra.F.id_mark 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_cons 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_recip 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_sqr 
                                               (x1::nil))::nil))::
                                               (algebra.Alg.Term 
                                               algebra.F.id_terms 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_s 
                                               (x1::nil))::nil))::nil))::nil)) 
           (algebra.Alg.Term algebra.F.id_active (x8::nil))
       
        (* <active(sqr(s(X_))),mark(s(add(sqr(X_),dbl(X_))))> *)
       | DP_R_xml_0_scc_38_large_large_large_2 :
        forall x8 x2, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term 
            algebra.F.id_s (x2::nil))::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_large (algebra.Alg.Term 
                                               algebra.F.id_mark 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_s 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_add 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_sqr (x2::nil))::
                                               (algebra.Alg.Term 
                                               algebra.F.id_dbl 
                                               (x2::nil))::nil))::nil))::nil)) 
           (algebra.Alg.Term algebra.F.id_active (x8::nil))
       
        (* <mark(recip(X_)),mark(X_)> *)
       | DP_R_xml_0_scc_38_large_large_large_3 :
        forall x8 x2, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_recip (x2::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_large (algebra.Alg.Term 
                                               algebra.F.id_mark (x2::nil)) 
           (algebra.Alg.Term algebra.F.id_mark (x8::nil))
       
        (* <mark(sqr(X_)),active(sqr(mark(X_)))> *)
       | DP_R_xml_0_scc_38_large_large_large_4 :
        forall x8 x2, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_large (algebra.Alg.Term 
                                               algebra.F.id_active 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_sqr 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_mark 
                                               (x2::nil))::nil))::nil)) 
           (algebra.Alg.Term algebra.F.id_mark (x8::nil))
       
        (* <mark(sqr(X_)),mark(X_)> *)
       | DP_R_xml_0_scc_38_large_large_large_5 :
        forall x8 x2, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_large (algebra.Alg.Term 
                                               algebra.F.id_mark (x2::nil)) 
           (algebra.Alg.Term algebra.F.id_mark (x8::nil))
       
        (* <mark(add(X1_,X2_)),active(add(mark(X1_),mark(X2_)))> *)
       | DP_R_xml_0_scc_38_large_large_large_6 :
        forall x8 x6 x5, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_large (algebra.Alg.Term 
                                               algebra.F.id_active 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_add 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_mark (x5::nil))::
                                               (algebra.Alg.Term 
                                               algebra.F.id_mark 
                                               (x6::nil))::nil))::nil)) 
           (algebra.Alg.Term algebra.F.id_mark (x8::nil))
       
        (* <mark(dbl(X_)),active(dbl(mark(X_)))> *)
       | DP_R_xml_0_scc_38_large_large_large_7 :
        forall x8 x2, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_large (algebra.Alg.Term 
                                               algebra.F.id_active 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_dbl 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_mark 
                                               (x2::nil))::nil))::nil)) 
           (algebra.Alg.Term algebra.F.id_mark (x8::nil))
       
        (* <mark(first(X1_,X2_)),active(first(mark(X1_),mark(X2_)))> *)
       | DP_R_xml_0_scc_38_large_large_large_8 :
        forall x8 x6 x5, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_large (algebra.Alg.Term 
                                               algebra.F.id_active 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_first 
                                               ((algebra.Alg.Term 
                                               algebra.F.id_mark (x5::nil))::
                                               (algebra.Alg.Term 
                                               algebra.F.id_mark 
                                               (x6::nil))::nil))::nil)) 
           (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     .
     
     
     Inductive DP_R_xml_0_scc_38_large_large_strict  :
      algebra.Alg.term ->algebra.Alg.term ->Prop := 
        (* <mark(terms(X_)),mark(X_)> *)
       | DP_R_xml_0_scc_38_large_large_strict_0 :
        forall x8 x2, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_terms (x2::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_strict (algebra.Alg.Term 
                                                algebra.F.id_mark (x2::nil)) 
           (algebra.Alg.Term algebra.F.id_mark (x8::nil))
       
        (* <mark(cons(X1_,X2_)),mark(X1_)> *)
       | DP_R_xml_0_scc_38_large_large_strict_1 :
        forall x8 x6 x5, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_strict (algebra.Alg.Term 
                                                algebra.F.id_mark (x5::nil)) 
           (algebra.Alg.Term algebra.F.id_mark (x8::nil))
       
        (* <active(dbl(s(X_))),mark(s(s(dbl(X_))))> *)
       | DP_R_xml_0_scc_38_large_large_strict_2 :
        forall x8 x2, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_dbl ((algebra.Alg.Term 
            algebra.F.id_s (x2::nil))::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_strict (algebra.Alg.Term 
                                                algebra.F.id_mark 
                                                ((algebra.Alg.Term 
                                                algebra.F.id_s 
                                                ((algebra.Alg.Term 
                                                algebra.F.id_s 
                                                ((algebra.Alg.Term 
                                                algebra.F.id_dbl 
                                                (x2::nil))::nil))::nil))::nil))
                                                 
           (algebra.Alg.Term algebra.F.id_active (x8::nil))
       
        (* <active(add(0,X_)),mark(X_)> *)
       | DP_R_xml_0_scc_38_large_large_strict_3 :
        forall x8 x2, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term 
            algebra.F.id_0 nil)::x2::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_strict (algebra.Alg.Term 
                                                algebra.F.id_mark (x2::nil)) 
           (algebra.Alg.Term algebra.F.id_active (x8::nil))
       
        (* <active(add(s(X_),Y_)),mark(s(add(X_,Y_)))> *)
       | DP_R_xml_0_scc_38_large_large_strict_4 :
        forall x8 x2 x3, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_add ((algebra.Alg.Term 
            algebra.F.id_s (x2::nil))::x3::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_strict (algebra.Alg.Term 
                                                algebra.F.id_mark 
                                                ((algebra.Alg.Term 
                                                algebra.F.id_s 
                                                ((algebra.Alg.Term 
                                                algebra.F.id_add (x2::
                                                x3::nil))::nil))::nil)) 
           (algebra.Alg.Term algebra.F.id_active (x8::nil))
       
        (* <active(first(s(X_),cons(Y_,Z_))),mark(cons(Y_,first(X_,Z_)))> *)
       | DP_R_xml_0_scc_38_large_large_strict_5 :
        forall x8 x4 x2 x3, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_first ((algebra.Alg.Term 
            algebra.F.id_s (x2::nil))::(algebra.Alg.Term algebra.F.id_cons 
            (x3::x4::nil))::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_strict (algebra.Alg.Term 
                                                algebra.F.id_mark 
                                                ((algebra.Alg.Term 
                                                algebra.F.id_cons (x3::
                                                (algebra.Alg.Term 
                                                algebra.F.id_first (x2::
                                                x4::nil))::nil))::nil)) 
           (algebra.Alg.Term algebra.F.id_active (x8::nil))
       
        (* <mark(add(X1_,X2_)),mark(X1_)> *)
       | DP_R_xml_0_scc_38_large_large_strict_6 :
        forall x8 x6 x5, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_strict (algebra.Alg.Term 
                                                algebra.F.id_mark (x5::nil)) 
           (algebra.Alg.Term algebra.F.id_mark (x8::nil))
       
        (* <mark(add(X1_,X2_)),mark(X2_)> *)
       | DP_R_xml_0_scc_38_large_large_strict_7 :
        forall x8 x6 x5, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_strict (algebra.Alg.Term 
                                                algebra.F.id_mark (x6::nil)) 
           (algebra.Alg.Term algebra.F.id_mark (x8::nil))
       
        (* <mark(dbl(X_)),mark(X_)> *)
       | DP_R_xml_0_scc_38_large_large_strict_8 :
        forall x8 x2, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_strict (algebra.Alg.Term 
                                                algebra.F.id_mark (x2::nil)) 
           (algebra.Alg.Term algebra.F.id_mark (x8::nil))
       
        (* <mark(first(X1_,X2_)),mark(X1_)> *)
       | DP_R_xml_0_scc_38_large_large_strict_9 :
        forall x8 x6 x5, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_strict (algebra.Alg.Term 
                                                algebra.F.id_mark (x5::nil)) 
           (algebra.Alg.Term algebra.F.id_mark (x8::nil))
       
        (* <mark(first(X1_,X2_)),mark(X2_)> *)
       | DP_R_xml_0_scc_38_large_large_strict_10 :
        forall x8 x6 x5, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
          DP_R_xml_0_scc_38_large_large_strict (algebra.Alg.Term 
                                                algebra.F.id_mark (x6::nil)) 
           (algebra.Alg.Term algebra.F.id_mark (x8::nil))
     .
     
     
     Module WF_DP_R_xml_0_scc_38_large_large_large.
      Inductive DP_R_xml_0_scc_38_large_large_large_large  :
       algebra.Alg.term ->algebra.Alg.term ->Prop := 
         (* <mark(terms(X_)),active(terms(mark(X_)))> *)
        | DP_R_xml_0_scc_38_large_large_large_large_0 :
         forall x8 x2, 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     
            (algebra.Alg.Term algebra.F.id_terms (x2::nil)) x8) ->
           DP_R_xml_0_scc_38_large_large_large_large (algebra.Alg.Term 
                                                      algebra.F.id_active 
                                                      ((algebra.Alg.Term 
                                                      algebra.F.id_terms 
                                                      ((algebra.Alg.Term 
                                                      algebra.F.id_mark 
                                                      (x2::nil))::nil))::nil)) 
            (algebra.Alg.Term algebra.F.id_mark (x8::nil))
        
         (* <active(terms(N_)),mark(cons(recip(sqr(N_)),terms(s(N_))))> *)
        | DP_R_xml_0_scc_38_large_large_large_large_1 :
         forall x8 x1, 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     
            (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
           DP_R_xml_0_scc_38_large_large_large_large (algebra.Alg.Term 
                                                      algebra.F.id_mark 
                                                      ((algebra.Alg.Term 
                                                      algebra.F.id_cons 
                                                      ((algebra.Alg.Term 
                                                      algebra.F.id_recip 
                                                      ((algebra.Alg.Term 
                                                      algebra.F.id_sqr 
                                                      (x1::nil))::nil))::
                                                      (algebra.Alg.Term 
                                                      algebra.F.id_terms 
                                                      ((algebra.Alg.Term 
                                                      algebra.F.id_s 
                                                      (x1::nil))::nil))::nil))::nil))
                                                       
            (algebra.Alg.Term algebra.F.id_active (x8::nil))
        
         (* <active(sqr(s(X_))),mark(s(add(sqr(X_),dbl(X_))))> *)
        | DP_R_xml_0_scc_38_large_large_large_large_2 :
         forall x8 x2, 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     
            (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term 
             algebra.F.id_s (x2::nil))::nil)) x8) ->
           DP_R_xml_0_scc_38_large_large_large_large (algebra.Alg.Term 
                                                      algebra.F.id_mark 
                                                      ((algebra.Alg.Term 
                                                      algebra.F.id_s 
                                                      ((algebra.Alg.Term 
                                                      algebra.F.id_add 
                                                      ((algebra.Alg.Term 
                                                      algebra.F.id_sqr 
                                                      (x2::nil))::
                                                      (algebra.Alg.Term 
                                                      algebra.F.id_dbl 
                                                      (x2::nil))::nil))::nil))::nil))
                                                       
            (algebra.Alg.Term algebra.F.id_active (x8::nil))
        
         (* <mark(recip(X_)),mark(X_)> *)
        | DP_R_xml_0_scc_38_large_large_large_large_3 :
         forall x8 x2, 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     
            (algebra.Alg.Term algebra.F.id_recip (x2::nil)) x8) ->
           DP_R_xml_0_scc_38_large_large_large_large (algebra.Alg.Term 
                                                      algebra.F.id_mark 
                                                      (x2::nil)) 
            (algebra.Alg.Term algebra.F.id_mark (x8::nil))
        
         (* <mark(sqr(X_)),active(sqr(mark(X_)))> *)
        | DP_R_xml_0_scc_38_large_large_large_large_4 :
         forall x8 x2, 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     
            (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) x8) ->
           DP_R_xml_0_scc_38_large_large_large_large (algebra.Alg.Term 
                                                      algebra.F.id_active 
                                                      ((algebra.Alg.Term 
                                                      algebra.F.id_sqr 
                                                      ((algebra.Alg.Term 
                                                      algebra.F.id_mark 
                                                      (x2::nil))::nil))::nil)) 
            (algebra.Alg.Term algebra.F.id_mark (x8::nil))
        
         (* <mark(sqr(X_)),mark(X_)> *)
        | DP_R_xml_0_scc_38_large_large_large_large_5 :
         forall x8 x2, 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     
            (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) x8) ->
           DP_R_xml_0_scc_38_large_large_large_large (algebra.Alg.Term 
                                                      algebra.F.id_mark 
                                                      (x2::nil)) 
            (algebra.Alg.Term algebra.F.id_mark (x8::nil))
        
         (* <mark(dbl(X_)),active(dbl(mark(X_)))> *)
        | DP_R_xml_0_scc_38_large_large_large_large_6 :
         forall x8 x2, 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     
            (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) x8) ->
           DP_R_xml_0_scc_38_large_large_large_large (algebra.Alg.Term 
                                                      algebra.F.id_active 
                                                      ((algebra.Alg.Term 
                                                      algebra.F.id_dbl 
                                                      ((algebra.Alg.Term 
                                                      algebra.F.id_mark 
                                                      (x2::nil))::nil))::nil)) 
            (algebra.Alg.Term algebra.F.id_mark (x8::nil))
        
         (* <mark(first(X1_,X2_)),active(first(mark(X1_),mark(X2_)))> *)
        | DP_R_xml_0_scc_38_large_large_large_large_7 :
         forall x8 x6 x5, 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     
            (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
           DP_R_xml_0_scc_38_large_large_large_large (algebra.Alg.Term 
                                                      algebra.F.id_active 
                                                      ((algebra.Alg.Term 
                                                      algebra.F.id_first 
                                                      ((algebra.Alg.Term 
                                                      algebra.F.id_mark 
                                                      (x5::nil))::
                                                      (algebra.Alg.Term 
                                                      algebra.F.id_mark 
                                                      (x6::nil))::nil))::nil)) 
            (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      .
      
      
      Inductive DP_R_xml_0_scc_38_large_large_large_strict  :
       algebra.Alg.term ->algebra.Alg.term ->Prop := 
         (* <mark(add(X1_,X2_)),active(add(mark(X1_),mark(X2_)))> *)
        | DP_R_xml_0_scc_38_large_large_large_strict_0 :
         forall x8 x6 x5, 
          (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                     
            (algebra.Alg.Term algebra.F.id_add (x5::x6::nil)) x8) ->
           DP_R_xml_0_scc_38_large_large_large_strict (algebra.Alg.Term 
                                                       algebra.F.id_active 
                                                       ((algebra.Alg.Term 
                                                       algebra.F.id_add 
                                                       ((algebra.Alg.Term 
                                                       algebra.F.id_mark 
                                                       (x5::nil))::
                                                       (algebra.Alg.Term 
                                                       algebra.F.id_mark 
                                                       (x6::nil))::nil))::nil))
                                                        
            (algebra.Alg.Term algebra.F.id_mark (x8::nil))
      .
      
      
      Module WF_DP_R_xml_0_scc_38_large_large_large_large.
       Inductive DP_R_xml_0_scc_38_large_large_large_large_large  :
        algebra.Alg.term ->algebra.Alg.term ->Prop := 
          (* <mark(recip(X_)),mark(X_)> *)
         | DP_R_xml_0_scc_38_large_large_large_large_large_0 :
          forall x8 x2, 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                      
             (algebra.Alg.Term algebra.F.id_recip (x2::nil)) x8) ->
            DP_R_xml_0_scc_38_large_large_large_large_large (algebra.Alg.Term 
                                                             algebra.F.id_mark 
                                                             (x2::nil)) 
             (algebra.Alg.Term algebra.F.id_mark (x8::nil))
       .
       
       
       Inductive DP_R_xml_0_scc_38_large_large_large_large_strict  :
        algebra.Alg.term ->algebra.Alg.term ->Prop := 
          (* <mark(terms(X_)),active(terms(mark(X_)))> *)
         | DP_R_xml_0_scc_38_large_large_large_large_strict_0 :
          forall x8 x2, 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                      
             (algebra.Alg.Term algebra.F.id_terms (x2::nil)) x8) ->
            DP_R_xml_0_scc_38_large_large_large_large_strict (algebra.Alg.Term 
                                                              algebra.F.id_active 
                                                              ((algebra.Alg.Term 
                                                              algebra.F.id_terms 
                                                              ((algebra.Alg.Term 
                                                              algebra.F.id_mark 
                                                              (x2::nil))::nil))::nil))
                                                               
             (algebra.Alg.Term algebra.F.id_mark (x8::nil))
         
          (* <active(terms(N_)),mark(cons(recip(sqr(N_)),terms(s(N_))))> *)
         | DP_R_xml_0_scc_38_large_large_large_large_strict_1 :
          forall x8 x1, 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                      
             (algebra.Alg.Term algebra.F.id_terms (x1::nil)) x8) ->
            DP_R_xml_0_scc_38_large_large_large_large_strict (algebra.Alg.Term 
                                                              algebra.F.id_mark 
                                                              ((algebra.Alg.Term 
                                                              algebra.F.id_cons 
                                                              ((algebra.Alg.Term 
                                                              algebra.F.id_recip 
                                                              ((algebra.Alg.Term 
                                                              algebra.F.id_sqr 
                                                              (x1::nil))::nil))::
                                                              
                                                              (algebra.Alg.Term 
                                                              algebra.F.id_terms 
                                                              ((algebra.Alg.Term 
                                                              algebra.F.id_s 
                                                              (x1::nil))::nil))::nil))::nil))
                                                               
             (algebra.Alg.Term algebra.F.id_active (x8::nil))
         
          (* <active(sqr(s(X_))),mark(s(add(sqr(X_),dbl(X_))))> *)
         | DP_R_xml_0_scc_38_large_large_large_large_strict_2 :
          forall x8 x2, 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                      
             (algebra.Alg.Term algebra.F.id_sqr ((algebra.Alg.Term 
              algebra.F.id_s (x2::nil))::nil)) x8) ->
            DP_R_xml_0_scc_38_large_large_large_large_strict (algebra.Alg.Term 
                                                              algebra.F.id_mark 
                                                              ((algebra.Alg.Term 
                                                              algebra.F.id_s 
                                                              ((algebra.Alg.Term 
                                                              algebra.F.id_add 
                                                              ((algebra.Alg.Term 
                                                              algebra.F.id_sqr 
                                                              (x2::nil))::
                                                              (algebra.Alg.Term 
                                                              algebra.F.id_dbl 
                                                              (x2::nil))::nil))::nil))::nil))
                                                               
             (algebra.Alg.Term algebra.F.id_active (x8::nil))
         
          (* <mark(sqr(X_)),active(sqr(mark(X_)))> *)
         | DP_R_xml_0_scc_38_large_large_large_large_strict_3 :
          forall x8 x2, 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                      
             (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) x8) ->
            DP_R_xml_0_scc_38_large_large_large_large_strict (algebra.Alg.Term 
                                                              algebra.F.id_active 
                                                              ((algebra.Alg.Term 
                                                              algebra.F.id_sqr 
                                                              ((algebra.Alg.Term 
                                                              algebra.F.id_mark 
                                                              (x2::nil))::nil))::nil))
                                                               
             (algebra.Alg.Term algebra.F.id_mark (x8::nil))
         
          (* <mark(sqr(X_)),mark(X_)> *)
         | DP_R_xml_0_scc_38_large_large_large_large_strict_4 :
          forall x8 x2, 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                      
             (algebra.Alg.Term algebra.F.id_sqr (x2::nil)) x8) ->
            DP_R_xml_0_scc_38_large_large_large_large_strict (algebra.Alg.Term 
                                                              algebra.F.id_mark 
                                                              (x2::nil)) 
             (algebra.Alg.Term algebra.F.id_mark (x8::nil))
         
          (* <mark(dbl(X_)),active(dbl(mark(X_)))> *)
         | DP_R_xml_0_scc_38_large_large_large_large_strict_5 :
          forall x8 x2, 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                      
             (algebra.Alg.Term algebra.F.id_dbl (x2::nil)) x8) ->
            DP_R_xml_0_scc_38_large_large_large_large_strict (algebra.Alg.Term 
                                                              algebra.F.id_active 
                                                              ((algebra.Alg.Term 
                                                              algebra.F.id_dbl 
                                                              ((algebra.Alg.Term 
                                                              algebra.F.id_mark 
                                                              (x2::nil))::nil))::nil))
                                                               
             (algebra.Alg.Term algebra.F.id_mark (x8::nil))
         
          (* <mark(first(X1_,X2_)),active(first(mark(X1_),mark(X2_)))> *)
         | DP_R_xml_0_scc_38_large_large_large_large_strict_6 :
          forall x8 x6 x5, 
           (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                      
             (algebra.Alg.Term algebra.F.id_first (x5::x6::nil)) x8) ->
            DP_R_xml_0_scc_38_large_large_large_large_strict (algebra.Alg.Term 
                                                              algebra.F.id_active 
                                                              ((algebra.Alg.Term 
                                                              algebra.F.id_first 
                                                              ((algebra.Alg.Term 
                                                              algebra.F.id_mark 
                                                              (x5::nil))::
                                                              (algebra.Alg.Term 
                                                              algebra.F.id_mark 
                                                              (x6::nil))::nil))::nil))
                                                               
             (algebra.Alg.Term algebra.F.id_mark (x8::nil))
       .
       
       
       Module WF_DP_R_xml_0_scc_38_large_large_large_large_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_active (x8:Z) := 1* x8.
        
        Definition P_id_add (x8:Z) (x9:Z) := 2* x9.
        
        Definition P_id_recip (x8:Z) := 1 + 2* x8.
        
        Definition P_id_mark (x8:Z) := 2* x8.
        
        Definition P_id_first (x8:Z) (x9:Z) := 0.
        
        Definition P_id_s (x8:Z) := 0.
        
        Definition P_id_terms (x8:Z) := 0.
        
        Definition P_id_dbl (x8:Z) := 2* x8.
        
        Definition P_id_sqr (x8:Z) := 2.
        
        Definition P_id_cons (x8:Z) (x9:Z) := 0.
        
        Definition P_id_nil  := 0.
        
        Definition P_id_0  := 0.
        
        Lemma P_id_active_monotonic :
         forall x8 x9, 
          (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
        Proof.
          intros x9 x8.
          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_add_monotonic :
         forall x8 x10 x9 x11, 
          (0 <= x11)/\ (x11 <= x10) ->
           (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
        Proof.
          intros x11 x10 x9 x8.
          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_recip_monotonic :
         forall x8 x9, 
          (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
        Proof.
          intros x9 x8.
          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_mark_monotonic :
         forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
        Proof.
          intros x9 x8.
          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_first_monotonic :
         forall x8 x10 x9 x11, 
          (0 <= x11)/\ (x11 <= x10) ->
           (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
        Proof.
          intros x11 x10 x9 x8.
          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_s_monotonic :
         forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
        Proof.
          intros x9 x8.
          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_terms_monotonic :
         forall x8 x9, 
          (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
        Proof.
          intros x9 x8.
          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_dbl_monotonic :
         forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
        Proof.
          intros x9 x8.
          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_sqr_monotonic :
         forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
        Proof.
          intros x9 x8.
          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_cons_monotonic :
         forall x8 x10 x9 x11, 
          (0 <= x11)/\ (x11 <= x10) ->
           (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
        Proof.
          intros x11 x10 x9 x8.
          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_active_bounded :
         forall x8, (0 <= x8) ->0 <= P_id_active x8.
        Proof.
          intros .
          
          cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
           (auto with zarith)||(repeat (translate_vars );prove_ineq ).
        Qed.
        
        Lemma P_id_add_bounded :
         forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
        Proof.
          intros .
          
          cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
           (auto with zarith)||(repeat (translate_vars );prove_ineq ).
        Qed.
        
        Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
        Proof.
          intros .
          
          cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
           (auto with zarith)||(repeat (translate_vars );prove_ineq ).
        Qed.
        
        Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
        Proof.
          intros .
          
          cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
           (auto with zarith)||(repeat (translate_vars );prove_ineq ).
        Qed.
        
        Lemma P_id_first_bounded :
         forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
        Proof.
          intros .
          
          cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
           (auto with zarith)||(repeat (translate_vars );prove_ineq ).
        Qed.
        
        Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
        Proof.
          intros .
          
          cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
           (auto with zarith)||(repeat (translate_vars );prove_ineq ).
        Qed.
        
        Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
        Proof.
          intros .
          
          cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
           (auto with zarith)||(repeat (translate_vars );prove_ineq ).
        Qed.
        
        Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
        Proof.
          intros .
          
          cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
           (auto with zarith)||(repeat (translate_vars );prove_ineq ).
        Qed.
        
        Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
        Proof.
          intros .
          
          cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
           (auto with zarith)||(repeat (translate_vars );prove_ineq ).
        Qed.
        
        Lemma P_id_cons_bounded :
         forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
        Proof.
          intros .
          
          cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
           (auto with zarith)||(repeat (translate_vars );prove_ineq ).
        Qed.
        
        Lemma P_id_nil_bounded : 0 <= P_id_nil .
        Proof.
          intros .
          
          cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
           (auto with zarith)||(repeat (translate_vars );prove_ineq ).
        Qed.
        
        Lemma P_id_0_bounded : 0 <= P_id_0 .
        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_active P_id_add P_id_recip P_id_mark 
           P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
           P_id_0.
        
        Lemma measure_equation :
         forall t, 
          measure t = match t with
                        | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                         P_id_active (measure x8)
                        | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                         P_id_add (measure x9) (measure x8)
                        | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                         P_id_recip (measure x8)
                        | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                         P_id_mark (measure x8)
                        | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                         P_id_first (measure x9) (measure x8)
                        | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                         P_id_s (measure x8)
                        | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                         P_id_terms (measure x8)
                        | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                         P_id_dbl (measure x8)
                        | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                         P_id_sqr (measure x8)
                        | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                         P_id_cons (measure x9) (measure x8)
                        | (algebra.Alg.Term algebra.F.id_nil nil) =>
                         P_id_nil 
                        | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                        | _ => 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_active_monotonic;assumption.
          intros ;apply P_id_add_monotonic;assumption.
          intros ;apply P_id_recip_monotonic;assumption.
          intros ;apply P_id_mark_monotonic;assumption.
          intros ;apply P_id_first_monotonic;assumption.
          intros ;apply P_id_s_monotonic;assumption.
          intros ;apply P_id_terms_monotonic;assumption.
          intros ;apply P_id_dbl_monotonic;assumption.
          intros ;apply P_id_sqr_monotonic;assumption.
          intros ;apply P_id_cons_monotonic;assumption.
          intros ;apply P_id_active_bounded;assumption.
          intros ;apply P_id_add_bounded;assumption.
          intros ;apply P_id_recip_bounded;assumption.
          intros ;apply P_id_mark_bounded;assumption.
          intros ;apply P_id_first_bounded;assumption.
          intros ;apply P_id_s_bounded;assumption.
          intros ;apply P_id_terms_bounded;assumption.
          intros ;apply P_id_dbl_bounded;assumption.
          intros ;apply P_id_sqr_bounded;assumption.
          intros ;apply P_id_cons_bounded;assumption.
          intros ;apply P_id_nil_bounded;assumption.
          intros ;apply P_id_0_bounded;assumption.
          apply rules_monotonic.
        Qed.
        
        Definition P_id_SQR (x8:Z) := 0.
        
        Definition P_id_DBL (x8:Z) := 0.
        
        Definition P_id_ACTIVE (x8:Z) := 0.
        
        Definition P_id_S (x8:Z) := 0.
        
        Definition P_id_CONS (x8:Z) (x9:Z) := 0.
        
        Definition P_id_TERMS (x8:Z) := 0.
        
        Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
        
        Definition P_id_MARK (x8:Z) := 2* x8.
        
        Definition P_id_ADD (x8:Z) (x9:Z) := 0.
        
        Definition P_id_RECIP (x8:Z) := 0.
        
        Lemma P_id_SQR_monotonic :
         forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
        Proof.
          intros x9 x8.
          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_DBL_monotonic :
         forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
        Proof.
          intros x9 x8.
          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_ACTIVE_monotonic :
         forall x8 x9, 
          (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
        Proof.
          intros x9 x8.
          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_S_monotonic :
         forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
        Proof.
          intros x9 x8.
          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_CONS_monotonic :
         forall x8 x10 x9 x11, 
          (0 <= x11)/\ (x11 <= x10) ->
           (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
        Proof.
          intros x11 x10 x9 x8.
          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_TERMS_monotonic :
         forall x8 x9, 
          (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
        Proof.
          intros x9 x8.
          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_FIRST_monotonic :
         forall x8 x10 x9 x11, 
          (0 <= x11)/\ (x11 <= x10) ->
           (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
        Proof.
          intros x11 x10 x9 x8.
          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_MARK_monotonic :
         forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
        Proof.
          intros x9 x8.
          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_ADD_monotonic :
         forall x8 x10 x9 x11, 
          (0 <= x11)/\ (x11 <= x10) ->
           (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
        Proof.
          intros x11 x10 x9 x8.
          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_RECIP_monotonic :
         forall x8 x9, 
          (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
        Proof.
          intros x9 x8.
          intros [H_1 H_0].
          
          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_active P_id_add P_id_recip P_id_mark 
           P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
           P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
           P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
        
        Lemma marked_measure_equation :
         forall t, 
          marked_measure t = match t with
                               | (algebra.Alg.Term algebra.F.id_sqr 
                                  (x8::nil)) =>
                                P_id_SQR (measure x8)
                               | (algebra.Alg.Term algebra.F.id_dbl 
                                  (x8::nil)) =>
                                P_id_DBL (measure x8)
                               | (algebra.Alg.Term algebra.F.id_active 
                                  (x8::nil)) =>
                                P_id_ACTIVE (measure x8)
                               | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                                P_id_S (measure x8)
                               | (algebra.Alg.Term algebra.F.id_cons (x9::
                                  x8::nil)) =>
                                P_id_CONS (measure x9) (measure x8)
                               | (algebra.Alg.Term algebra.F.id_terms 
                                  (x8::nil)) =>
                                P_id_TERMS (measure x8)
                               | (algebra.Alg.Term algebra.F.id_first (x9::
                                  x8::nil)) =>
                                P_id_FIRST (measure x9) (measure x8)
                               | (algebra.Alg.Term algebra.F.id_mark 
                                  (x8::nil)) =>
                                P_id_MARK (measure x8)
                               | (algebra.Alg.Term algebra.F.id_add (x9::
                                  x8::nil)) =>
                                P_id_ADD (measure x9) (measure x8)
                               | (algebra.Alg.Term algebra.F.id_recip 
                                  (x8::nil)) =>
                                P_id_RECIP (measure x8)
                               | _ => 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_active_monotonic;assumption.
          intros ;apply P_id_add_monotonic;assumption.
          intros ;apply P_id_recip_monotonic;assumption.
          intros ;apply P_id_mark_monotonic;assumption.
          intros ;apply P_id_first_monotonic;assumption.
          intros ;apply P_id_s_monotonic;assumption.
          intros ;apply P_id_terms_monotonic;assumption.
          intros ;apply P_id_dbl_monotonic;assumption.
          intros ;apply P_id_sqr_monotonic;assumption.
          intros ;apply P_id_cons_monotonic;assumption.
          intros ;apply P_id_active_bounded;assumption.
          intros ;apply P_id_add_bounded;assumption.
          intros ;apply P_id_recip_bounded;assumption.
          intros ;apply P_id_mark_bounded;assumption.
          intros ;apply P_id_first_bounded;assumption.
          intros ;apply P_id_s_bounded;assumption.
          intros ;apply P_id_terms_bounded;assumption.
          intros ;apply P_id_dbl_bounded;assumption.
          intros ;apply P_id_sqr_bounded;assumption.
          intros ;apply P_id_cons_bounded;assumption.
          intros ;apply P_id_nil_bounded;assumption.
          intros ;apply P_id_0_bounded;assumption.
          apply rules_monotonic.
          intros ;apply P_id_SQR_monotonic;assumption.
          intros ;apply P_id_DBL_monotonic;assumption.
          intros ;apply P_id_ACTIVE_monotonic;assumption.
          intros ;apply P_id_S_monotonic;assumption.
          intros ;apply P_id_CONS_monotonic;assumption.
          intros ;apply P_id_TERMS_monotonic;assumption.
          intros ;apply P_id_FIRST_monotonic;assumption.
          intros ;apply P_id_MARK_monotonic;assumption.
          intros ;apply P_id_ADD_monotonic;assumption.
          intros ;apply P_id_RECIP_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_38_large_large_large_large.DP_R_xml_0_scc_38_large_large_large_large_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_38_large_large_large_large_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_active (x8:Z) := 1* x8.
       
       Definition P_id_add (x8:Z) (x9:Z) := 1* x9.
       
       Definition P_id_recip (x8:Z) := 2* x8.
       
       Definition P_id_mark (x8:Z) := 1* x8.
       
       Definition P_id_first (x8:Z) (x9:Z) := 2.
       
       Definition P_id_s (x8:Z) := 0.
       
       Definition P_id_terms (x8:Z) := 2.
       
       Definition P_id_dbl (x8:Z) := 0.
       
       Definition P_id_sqr (x8:Z) := 2 + 2* x8.
       
       Definition P_id_cons (x8:Z) (x9:Z) := 0.
       
       Definition P_id_nil  := 0.
       
       Definition P_id_0  := 0.
       
       Lemma P_id_active_monotonic :
        forall x8 x9, 
         (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
       Proof.
         intros x9 x8.
         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_add_monotonic :
        forall x8 x10 x9 x11, 
         (0 <= x11)/\ (x11 <= x10) ->
          (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
       Proof.
         intros x11 x10 x9 x8.
         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_recip_monotonic :
        forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
       Proof.
         intros x9 x8.
         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_mark_monotonic :
        forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
       Proof.
         intros x9 x8.
         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_first_monotonic :
        forall x8 x10 x9 x11, 
         (0 <= x11)/\ (x11 <= x10) ->
          (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
       Proof.
         intros x11 x10 x9 x8.
         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_s_monotonic :
        forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
       Proof.
         intros x9 x8.
         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_terms_monotonic :
        forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
       Proof.
         intros x9 x8.
         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_dbl_monotonic :
        forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
       Proof.
         intros x9 x8.
         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_sqr_monotonic :
        forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
       Proof.
         intros x9 x8.
         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_cons_monotonic :
        forall x8 x10 x9 x11, 
         (0 <= x11)/\ (x11 <= x10) ->
          (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
       Proof.
         intros x11 x10 x9 x8.
         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_active_bounded :
        forall x8, (0 <= x8) ->0 <= P_id_active x8.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_add_bounded :
        forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_first_bounded :
        forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_cons_bounded :
        forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_nil_bounded : 0 <= P_id_nil .
       Proof.
         intros .
         
         cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
          (auto with zarith)||(repeat (translate_vars );prove_ineq ).
       Qed.
       
       Lemma P_id_0_bounded : 0 <= P_id_0 .
       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_active P_id_add P_id_recip P_id_mark 
          P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
          P_id_0.
       
       Lemma measure_equation :
        forall t, 
         measure t = match t with
                       | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                        P_id_active (measure x8)
                       | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                        P_id_add (measure x9) (measure x8)
                       | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                        P_id_recip (measure x8)
                       | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                        P_id_mark (measure x8)
                       | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                        P_id_first (measure x9) (measure x8)
                       | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                        P_id_s (measure x8)
                       | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                        P_id_terms (measure x8)
                       | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                        P_id_dbl (measure x8)
                       | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                        P_id_sqr (measure x8)
                       | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                        P_id_cons (measure x9) (measure x8)
                       | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                       | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                       | _ => 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_active_monotonic;assumption.
         intros ;apply P_id_add_monotonic;assumption.
         intros ;apply P_id_recip_monotonic;assumption.
         intros ;apply P_id_mark_monotonic;assumption.
         intros ;apply P_id_first_monotonic;assumption.
         intros ;apply P_id_s_monotonic;assumption.
         intros ;apply P_id_terms_monotonic;assumption.
         intros ;apply P_id_dbl_monotonic;assumption.
         intros ;apply P_id_sqr_monotonic;assumption.
         intros ;apply P_id_cons_monotonic;assumption.
         intros ;apply P_id_active_bounded;assumption.
         intros ;apply P_id_add_bounded;assumption.
         intros ;apply P_id_recip_bounded;assumption.
         intros ;apply P_id_mark_bounded;assumption.
         intros ;apply P_id_first_bounded;assumption.
         intros ;apply P_id_s_bounded;assumption.
         intros ;apply P_id_terms_bounded;assumption.
         intros ;apply P_id_dbl_bounded;assumption.
         intros ;apply P_id_sqr_bounded;assumption.
         intros ;apply P_id_cons_bounded;assumption.
         intros ;apply P_id_nil_bounded;assumption.
         intros ;apply P_id_0_bounded;assumption.
         apply rules_monotonic.
       Qed.
       
       Definition P_id_SQR (x8:Z) := 0.
       
       Definition P_id_DBL (x8:Z) := 0.
       
       Definition P_id_ACTIVE (x8:Z) := 1 + 2* x8.
       
       Definition P_id_S (x8:Z) := 0.
       
       Definition P_id_CONS (x8:Z) (x9:Z) := 0.
       
       Definition P_id_TERMS (x8:Z) := 0.
       
       Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
       
       Definition P_id_MARK (x8:Z) := 2 + 3* x8.
       
       Definition P_id_ADD (x8:Z) (x9:Z) := 0.
       
       Definition P_id_RECIP (x8:Z) := 0.
       
       Lemma P_id_SQR_monotonic :
        forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
       Proof.
         intros x9 x8.
         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_DBL_monotonic :
        forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
       Proof.
         intros x9 x8.
         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_ACTIVE_monotonic :
        forall x8 x9, 
         (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
       Proof.
         intros x9 x8.
         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_S_monotonic :
        forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
       Proof.
         intros x9 x8.
         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_CONS_monotonic :
        forall x8 x10 x9 x11, 
         (0 <= x11)/\ (x11 <= x10) ->
          (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
       Proof.
         intros x11 x10 x9 x8.
         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_TERMS_monotonic :
        forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
       Proof.
         intros x9 x8.
         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_FIRST_monotonic :
        forall x8 x10 x9 x11, 
         (0 <= x11)/\ (x11 <= x10) ->
          (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
       Proof.
         intros x11 x10 x9 x8.
         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_MARK_monotonic :
        forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
       Proof.
         intros x9 x8.
         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_ADD_monotonic :
        forall x8 x10 x9 x11, 
         (0 <= x11)/\ (x11 <= x10) ->
          (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
       Proof.
         intros x11 x10 x9 x8.
         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_RECIP_monotonic :
        forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
       Proof.
         intros x9 x8.
         intros [H_1 H_0].
         
         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_active P_id_add P_id_recip P_id_mark 
          P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
          P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
          P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
       
       Lemma marked_measure_equation :
        forall t, 
         marked_measure t = match t with
                              | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                               P_id_SQR (measure x8)
                              | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                               P_id_DBL (measure x8)
                              | (algebra.Alg.Term algebra.F.id_active 
                                 (x8::nil)) =>
                               P_id_ACTIVE (measure x8)
                              | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                               P_id_S (measure x8)
                              | (algebra.Alg.Term algebra.F.id_cons (x9::
                                 x8::nil)) =>
                               P_id_CONS (measure x9) (measure x8)
                              | (algebra.Alg.Term algebra.F.id_terms 
                                 (x8::nil)) =>
                               P_id_TERMS (measure x8)
                              | (algebra.Alg.Term algebra.F.id_first (x9::
                                 x8::nil)) =>
                               P_id_FIRST (measure x9) (measure x8)
                              | (algebra.Alg.Term algebra.F.id_mark 
                                 (x8::nil)) =>
                               P_id_MARK (measure x8)
                              | (algebra.Alg.Term algebra.F.id_add (x9::
                                 x8::nil)) =>
                               P_id_ADD (measure x9) (measure x8)
                              | (algebra.Alg.Term algebra.F.id_recip 
                                 (x8::nil)) =>
                               P_id_RECIP (measure x8)
                              | _ => 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_active_monotonic;assumption.
         intros ;apply P_id_add_monotonic;assumption.
         intros ;apply P_id_recip_monotonic;assumption.
         intros ;apply P_id_mark_monotonic;assumption.
         intros ;apply P_id_first_monotonic;assumption.
         intros ;apply P_id_s_monotonic;assumption.
         intros ;apply P_id_terms_monotonic;assumption.
         intros ;apply P_id_dbl_monotonic;assumption.
         intros ;apply P_id_sqr_monotonic;assumption.
         intros ;apply P_id_cons_monotonic;assumption.
         intros ;apply P_id_active_bounded;assumption.
         intros ;apply P_id_add_bounded;assumption.
         intros ;apply P_id_recip_bounded;assumption.
         intros ;apply P_id_mark_bounded;assumption.
         intros ;apply P_id_first_bounded;assumption.
         intros ;apply P_id_s_bounded;assumption.
         intros ;apply P_id_terms_bounded;assumption.
         intros ;apply P_id_dbl_bounded;assumption.
         intros ;apply P_id_sqr_bounded;assumption.
         intros ;apply P_id_cons_bounded;assumption.
         intros ;apply P_id_nil_bounded;assumption.
         intros ;apply P_id_0_bounded;assumption.
         apply rules_monotonic.
         intros ;apply P_id_SQR_monotonic;assumption.
         intros ;apply P_id_DBL_monotonic;assumption.
         intros ;apply P_id_ACTIVE_monotonic;assumption.
         intros ;apply P_id_S_monotonic;assumption.
         intros ;apply P_id_CONS_monotonic;assumption.
         intros ;apply P_id_TERMS_monotonic;assumption.
         intros ;apply P_id_FIRST_monotonic;assumption.
         intros ;apply P_id_MARK_monotonic;assumption.
         intros ;apply P_id_ADD_monotonic;assumption.
         intros ;apply P_id_RECIP_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_38_large_large_large_large_strict_in_lt :
        Relation_Definitions.inclusion _ 
         DP_R_xml_0_scc_38_large_large_large_large_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_38_large_large_large_large_large_in_le :
        Relation_Definitions.inclusion _ 
         DP_R_xml_0_scc_38_large_large_large_large_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_38_large_large_large_large_large  := 
         WF_DP_R_xml_0_scc_38_large_large_large_large_large.wf.
       
       
       Lemma wf :
        well_founded WF_DP_R_xml_0_scc_38_large_large_large.DP_R_xml_0_scc_38_large_large_large_large
         .
       Proof.
         intros x.
         apply (well_founded_ind wf_lt).
         clear x.
         intros x.
         pattern x.
         apply (@Acc_ind _ DP_R_xml_0_scc_38_large_large_large_large_large).
         clear x.
         intros x _ IHx IHx'.
         constructor.
         intros y H.
         
         destruct H;
          (apply IHx';
            apply DP_R_xml_0_scc_38_large_large_large_large_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_38_large_large_large_large_large_in_le;
             econstructor eassumption])).
         apply wf_DP_R_xml_0_scc_38_large_large_large_large_large.
       Qed.
      End WF_DP_R_xml_0_scc_38_large_large_large_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_active (x8:Z) := 0.
      
      Definition P_id_add (x8:Z) (x9:Z) := 0.
      
      Definition P_id_recip (x8:Z) := 0.
      
      Definition P_id_mark (x8:Z) := 0.
      
      Definition P_id_first (x8:Z) (x9:Z) := 1.
      
      Definition P_id_s (x8:Z) := 0.
      
      Definition P_id_terms (x8:Z) := 1.
      
      Definition P_id_dbl (x8:Z) := 1.
      
      Definition P_id_sqr (x8:Z) := 1.
      
      Definition P_id_cons (x8:Z) (x9:Z) := 0.
      
      Definition P_id_nil  := 0.
      
      Definition P_id_0  := 0.
      
      Lemma P_id_active_monotonic :
       forall x8 x9, 
        (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
      Proof.
        intros x9 x8.
        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_add_monotonic :
       forall x8 x10 x9 x11, 
        (0 <= x11)/\ (x11 <= x10) ->
         (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
      Proof.
        intros x11 x10 x9 x8.
        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_recip_monotonic :
       forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
      Proof.
        intros x9 x8.
        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_mark_monotonic :
       forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
      Proof.
        intros x9 x8.
        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_first_monotonic :
       forall x8 x10 x9 x11, 
        (0 <= x11)/\ (x11 <= x10) ->
         (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
      Proof.
        intros x11 x10 x9 x8.
        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_s_monotonic :
       forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
      Proof.
        intros x9 x8.
        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_terms_monotonic :
       forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
      Proof.
        intros x9 x8.
        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_dbl_monotonic :
       forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
      Proof.
        intros x9 x8.
        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_sqr_monotonic :
       forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
      Proof.
        intros x9 x8.
        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_cons_monotonic :
       forall x8 x10 x9 x11, 
        (0 <= x11)/\ (x11 <= x10) ->
         (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
      Proof.
        intros x11 x10 x9 x8.
        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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_add_bounded :
       forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_first_bounded :
       forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_cons_bounded :
       forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_nil_bounded : 0 <= P_id_nil .
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_0_bounded : 0 <= P_id_0 .
      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_active P_id_add P_id_recip P_id_mark 
         P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
         P_id_0.
      
      Lemma measure_equation :
       forall t, 
        measure t = match t with
                      | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                       P_id_active (measure x8)
                      | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                       P_id_add (measure x9) (measure x8)
                      | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                       P_id_recip (measure x8)
                      | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                       P_id_mark (measure x8)
                      | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                       P_id_first (measure x9) (measure x8)
                      | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                       P_id_s (measure x8)
                      | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                       P_id_terms (measure x8)
                      | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                       P_id_dbl (measure x8)
                      | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                       P_id_sqr (measure x8)
                      | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                       P_id_cons (measure x9) (measure x8)
                      | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                      | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                      | _ => 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_active_monotonic;assumption.
        intros ;apply P_id_add_monotonic;assumption.
        intros ;apply P_id_recip_monotonic;assumption.
        intros ;apply P_id_mark_monotonic;assumption.
        intros ;apply P_id_first_monotonic;assumption.
        intros ;apply P_id_s_monotonic;assumption.
        intros ;apply P_id_terms_monotonic;assumption.
        intros ;apply P_id_dbl_monotonic;assumption.
        intros ;apply P_id_sqr_monotonic;assumption.
        intros ;apply P_id_cons_monotonic;assumption.
        intros ;apply P_id_active_bounded;assumption.
        intros ;apply P_id_add_bounded;assumption.
        intros ;apply P_id_recip_bounded;assumption.
        intros ;apply P_id_mark_bounded;assumption.
        intros ;apply P_id_first_bounded;assumption.
        intros ;apply P_id_s_bounded;assumption.
        intros ;apply P_id_terms_bounded;assumption.
        intros ;apply P_id_dbl_bounded;assumption.
        intros ;apply P_id_sqr_bounded;assumption.
        intros ;apply P_id_cons_bounded;assumption.
        intros ;apply P_id_nil_bounded;assumption.
        intros ;apply P_id_0_bounded;assumption.
        apply rules_monotonic.
      Qed.
      
      Definition P_id_SQR (x8:Z) := 0.
      
      Definition P_id_DBL (x8:Z) := 0.
      
      Definition P_id_ACTIVE (x8:Z) := 1* x8.
      
      Definition P_id_S (x8:Z) := 0.
      
      Definition P_id_CONS (x8:Z) (x9:Z) := 0.
      
      Definition P_id_TERMS (x8:Z) := 0.
      
      Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
      
      Definition P_id_MARK (x8:Z) := 1.
      
      Definition P_id_ADD (x8:Z) (x9:Z) := 0.
      
      Definition P_id_RECIP (x8:Z) := 0.
      
      Lemma P_id_SQR_monotonic :
       forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
      Proof.
        intros x9 x8.
        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_DBL_monotonic :
       forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
      Proof.
        intros x9 x8.
        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_ACTIVE_monotonic :
       forall x8 x9, 
        (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
      Proof.
        intros x9 x8.
        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_S_monotonic :
       forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
      Proof.
        intros x9 x8.
        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_CONS_monotonic :
       forall x8 x10 x9 x11, 
        (0 <= x11)/\ (x11 <= x10) ->
         (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
      Proof.
        intros x11 x10 x9 x8.
        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_TERMS_monotonic :
       forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
      Proof.
        intros x9 x8.
        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_FIRST_monotonic :
       forall x8 x10 x9 x11, 
        (0 <= x11)/\ (x11 <= x10) ->
         (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
      Proof.
        intros x11 x10 x9 x8.
        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_MARK_monotonic :
       forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
      Proof.
        intros x9 x8.
        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_ADD_monotonic :
       forall x8 x10 x9 x11, 
        (0 <= x11)/\ (x11 <= x10) ->
         (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
      Proof.
        intros x11 x10 x9 x8.
        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_RECIP_monotonic :
       forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
      Proof.
        intros x9 x8.
        intros [H_1 H_0].
        
        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_active P_id_add P_id_recip P_id_mark 
         P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
         P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
         P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
      
      Lemma marked_measure_equation :
       forall t, 
        marked_measure t = match t with
                             | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                              P_id_SQR (measure x8)
                             | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                              P_id_DBL (measure x8)
                             | (algebra.Alg.Term algebra.F.id_active 
                                (x8::nil)) =>
                              P_id_ACTIVE (measure x8)
                             | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                              P_id_S (measure x8)
                             | (algebra.Alg.Term algebra.F.id_cons (x9::
                                x8::nil)) =>
                              P_id_CONS (measure x9) (measure x8)
                             | (algebra.Alg.Term algebra.F.id_terms 
                                (x8::nil)) =>
                              P_id_TERMS (measure x8)
                             | (algebra.Alg.Term algebra.F.id_first (x9::
                                x8::nil)) =>
                              P_id_FIRST (measure x9) (measure x8)
                             | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                              P_id_MARK (measure x8)
                             | (algebra.Alg.Term algebra.F.id_add (x9::
                                x8::nil)) =>
                              P_id_ADD (measure x9) (measure x8)
                             | (algebra.Alg.Term algebra.F.id_recip 
                                (x8::nil)) =>
                              P_id_RECIP (measure x8)
                             | _ => 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_active_monotonic;assumption.
        intros ;apply P_id_add_monotonic;assumption.
        intros ;apply P_id_recip_monotonic;assumption.
        intros ;apply P_id_mark_monotonic;assumption.
        intros ;apply P_id_first_monotonic;assumption.
        intros ;apply P_id_s_monotonic;assumption.
        intros ;apply P_id_terms_monotonic;assumption.
        intros ;apply P_id_dbl_monotonic;assumption.
        intros ;apply P_id_sqr_monotonic;assumption.
        intros ;apply P_id_cons_monotonic;assumption.
        intros ;apply P_id_active_bounded;assumption.
        intros ;apply P_id_add_bounded;assumption.
        intros ;apply P_id_recip_bounded;assumption.
        intros ;apply P_id_mark_bounded;assumption.
        intros ;apply P_id_first_bounded;assumption.
        intros ;apply P_id_s_bounded;assumption.
        intros ;apply P_id_terms_bounded;assumption.
        intros ;apply P_id_dbl_bounded;assumption.
        intros ;apply P_id_sqr_bounded;assumption.
        intros ;apply P_id_cons_bounded;assumption.
        intros ;apply P_id_nil_bounded;assumption.
        intros ;apply P_id_0_bounded;assumption.
        apply rules_monotonic.
        intros ;apply P_id_SQR_monotonic;assumption.
        intros ;apply P_id_DBL_monotonic;assumption.
        intros ;apply P_id_ACTIVE_monotonic;assumption.
        intros ;apply P_id_S_monotonic;assumption.
        intros ;apply P_id_CONS_monotonic;assumption.
        intros ;apply P_id_TERMS_monotonic;assumption.
        intros ;apply P_id_FIRST_monotonic;assumption.
        intros ;apply P_id_MARK_monotonic;assumption.
        intros ;apply P_id_ADD_monotonic;assumption.
        intros ;apply P_id_RECIP_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_38_large_large_large_strict_in_lt :
       Relation_Definitions.inclusion _ 
        DP_R_xml_0_scc_38_large_large_large_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_38_large_large_large_large_in_le :
       Relation_Definitions.inclusion _ 
        DP_R_xml_0_scc_38_large_large_large_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_38_large_large_large_large  := 
        WF_DP_R_xml_0_scc_38_large_large_large_large.wf.
      
      
      Lemma wf :
       well_founded WF_DP_R_xml_0_scc_38_large_large.DP_R_xml_0_scc_38_large_large_large
        .
      Proof.
        intros x.
        apply (well_founded_ind wf_lt).
        clear x.
        intros x.
        pattern x.
        apply (@Acc_ind _ DP_R_xml_0_scc_38_large_large_large_large).
        clear x.
        intros x _ IHx IHx'.
        constructor.
        intros y H.
        
        destruct H;
         (apply IHx';apply DP_R_xml_0_scc_38_large_large_large_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_38_large_large_large_large_in_le;
            econstructor eassumption])).
        apply wf_DP_R_xml_0_scc_38_large_large_large_large.
      Qed.
     End WF_DP_R_xml_0_scc_38_large_large_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_active (x8:Z) := 1* x8.
     
     Definition P_id_add (x8:Z) (x9:Z) := 2 + 1* x8 + 1* x9.
     
     Definition P_id_recip (x8:Z) := 1* x8.
     
     Definition P_id_mark (x8:Z) := 1* x8.
     
     Definition P_id_first (x8:Z) (x9:Z) := 1 + 3* x8 + 2* x9.
     
     Definition P_id_s (x8:Z) := 3.
     
     Definition P_id_terms (x8:Z) := 3 + 1* x8.
     
     Definition P_id_dbl (x8:Z) := 2 + 2* x8.
     
     Definition P_id_sqr (x8:Z) := 1* x8.
     
     Definition P_id_cons (x8:Z) (x9:Z) := 3 + 1* x8.
     
     Definition P_id_nil  := 1.
     
     Definition P_id_0  := 0.
     
     Lemma P_id_active_monotonic :
      forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
     Proof.
       intros x9 x8.
       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_add_monotonic :
      forall x8 x10 x9 x11, 
       (0 <= x11)/\ (x11 <= x10) ->
        (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
     Proof.
       intros x11 x10 x9 x8.
       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_recip_monotonic :
      forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
     Proof.
       intros x9 x8.
       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_mark_monotonic :
      forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
     Proof.
       intros x9 x8.
       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_first_monotonic :
      forall x8 x10 x9 x11, 
       (0 <= x11)/\ (x11 <= x10) ->
        (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
     Proof.
       intros x11 x10 x9 x8.
       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_s_monotonic :
      forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
     Proof.
       intros x9 x8.
       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_terms_monotonic :
      forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
     Proof.
       intros x9 x8.
       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_dbl_monotonic :
      forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
     Proof.
       intros x9 x8.
       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_sqr_monotonic :
      forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
     Proof.
       intros x9 x8.
       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_cons_monotonic :
      forall x8 x10 x9 x11, 
       (0 <= x11)/\ (x11 <= x10) ->
        (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
     Proof.
       intros x11 x10 x9 x8.
       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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_add_bounded :
      forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_first_bounded :
      forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_cons_bounded :
      forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_nil_bounded : 0 <= P_id_nil .
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_0_bounded : 0 <= P_id_0 .
     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_active P_id_add P_id_recip P_id_mark 
        P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
        P_id_0.
     
     Lemma measure_equation :
      forall t, 
       measure t = match t with
                     | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                      P_id_active (measure x8)
                     | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                      P_id_add (measure x9) (measure x8)
                     | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                      P_id_recip (measure x8)
                     | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                      P_id_mark (measure x8)
                     | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                      P_id_first (measure x9) (measure x8)
                     | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                      P_id_s (measure x8)
                     | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                      P_id_terms (measure x8)
                     | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                      P_id_dbl (measure x8)
                     | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                      P_id_sqr (measure x8)
                     | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                      P_id_cons (measure x9) (measure x8)
                     | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                     | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                     | _ => 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_active_monotonic;assumption.
       intros ;apply P_id_add_monotonic;assumption.
       intros ;apply P_id_recip_monotonic;assumption.
       intros ;apply P_id_mark_monotonic;assumption.
       intros ;apply P_id_first_monotonic;assumption.
       intros ;apply P_id_s_monotonic;assumption.
       intros ;apply P_id_terms_monotonic;assumption.
       intros ;apply P_id_dbl_monotonic;assumption.
       intros ;apply P_id_sqr_monotonic;assumption.
       intros ;apply P_id_cons_monotonic;assumption.
       intros ;apply P_id_active_bounded;assumption.
       intros ;apply P_id_add_bounded;assumption.
       intros ;apply P_id_recip_bounded;assumption.
       intros ;apply P_id_mark_bounded;assumption.
       intros ;apply P_id_first_bounded;assumption.
       intros ;apply P_id_s_bounded;assumption.
       intros ;apply P_id_terms_bounded;assumption.
       intros ;apply P_id_dbl_bounded;assumption.
       intros ;apply P_id_sqr_bounded;assumption.
       intros ;apply P_id_cons_bounded;assumption.
       intros ;apply P_id_nil_bounded;assumption.
       intros ;apply P_id_0_bounded;assumption.
       apply rules_monotonic.
     Qed.
     
     Definition P_id_SQR (x8:Z) := 0.
     
     Definition P_id_DBL (x8:Z) := 0.
     
     Definition P_id_ACTIVE (x8:Z) := 1* x8.
     
     Definition P_id_S (x8:Z) := 0.
     
     Definition P_id_CONS (x8:Z) (x9:Z) := 0.
     
     Definition P_id_TERMS (x8:Z) := 0.
     
     Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
     
     Definition P_id_MARK (x8:Z) := 1* x8.
     
     Definition P_id_ADD (x8:Z) (x9:Z) := 0.
     
     Definition P_id_RECIP (x8:Z) := 0.
     
     Lemma P_id_SQR_monotonic :
      forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
     Proof.
       intros x9 x8.
       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_DBL_monotonic :
      forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
     Proof.
       intros x9 x8.
       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_ACTIVE_monotonic :
      forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
     Proof.
       intros x9 x8.
       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_S_monotonic :
      forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
     Proof.
       intros x9 x8.
       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_CONS_monotonic :
      forall x8 x10 x9 x11, 
       (0 <= x11)/\ (x11 <= x10) ->
        (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
     Proof.
       intros x11 x10 x9 x8.
       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_TERMS_monotonic :
      forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
     Proof.
       intros x9 x8.
       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_FIRST_monotonic :
      forall x8 x10 x9 x11, 
       (0 <= x11)/\ (x11 <= x10) ->
        (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
     Proof.
       intros x11 x10 x9 x8.
       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_MARK_monotonic :
      forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
     Proof.
       intros x9 x8.
       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_ADD_monotonic :
      forall x8 x10 x9 x11, 
       (0 <= x11)/\ (x11 <= x10) ->
        (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
     Proof.
       intros x11 x10 x9 x8.
       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_RECIP_monotonic :
      forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
     Proof.
       intros x9 x8.
       intros [H_1 H_0].
       
       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_active P_id_add P_id_recip P_id_mark 
        P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
        P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
        P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
     
     Lemma marked_measure_equation :
      forall t, 
       marked_measure t = match t with
                            | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                             P_id_SQR (measure x8)
                            | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                             P_id_DBL (measure x8)
                            | (algebra.Alg.Term algebra.F.id_active 
                               (x8::nil)) =>
                             P_id_ACTIVE (measure x8)
                            | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                             P_id_S (measure x8)
                            | (algebra.Alg.Term algebra.F.id_cons (x9::
                               x8::nil)) =>
                             P_id_CONS (measure x9) (measure x8)
                            | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                             P_id_TERMS (measure x8)
                            | (algebra.Alg.Term algebra.F.id_first (x9::
                               x8::nil)) =>
                             P_id_FIRST (measure x9) (measure x8)
                            | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                             P_id_MARK (measure x8)
                            | (algebra.Alg.Term algebra.F.id_add (x9::
                               x8::nil)) =>
                             P_id_ADD (measure x9) (measure x8)
                            | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                             P_id_RECIP (measure x8)
                            | _ => 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_active_monotonic;assumption.
       intros ;apply P_id_add_monotonic;assumption.
       intros ;apply P_id_recip_monotonic;assumption.
       intros ;apply P_id_mark_monotonic;assumption.
       intros ;apply P_id_first_monotonic;assumption.
       intros ;apply P_id_s_monotonic;assumption.
       intros ;apply P_id_terms_monotonic;assumption.
       intros ;apply P_id_dbl_monotonic;assumption.
       intros ;apply P_id_sqr_monotonic;assumption.
       intros ;apply P_id_cons_monotonic;assumption.
       intros ;apply P_id_active_bounded;assumption.
       intros ;apply P_id_add_bounded;assumption.
       intros ;apply P_id_recip_bounded;assumption.
       intros ;apply P_id_mark_bounded;assumption.
       intros ;apply P_id_first_bounded;assumption.
       intros ;apply P_id_s_bounded;assumption.
       intros ;apply P_id_terms_bounded;assumption.
       intros ;apply P_id_dbl_bounded;assumption.
       intros ;apply P_id_sqr_bounded;assumption.
       intros ;apply P_id_cons_bounded;assumption.
       intros ;apply P_id_nil_bounded;assumption.
       intros ;apply P_id_0_bounded;assumption.
       apply rules_monotonic.
       intros ;apply P_id_SQR_monotonic;assumption.
       intros ;apply P_id_DBL_monotonic;assumption.
       intros ;apply P_id_ACTIVE_monotonic;assumption.
       intros ;apply P_id_S_monotonic;assumption.
       intros ;apply P_id_CONS_monotonic;assumption.
       intros ;apply P_id_TERMS_monotonic;assumption.
       intros ;apply P_id_FIRST_monotonic;assumption.
       intros ;apply P_id_MARK_monotonic;assumption.
       intros ;apply P_id_ADD_monotonic;assumption.
       intros ;apply P_id_RECIP_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_38_large_large_strict_in_lt :
      Relation_Definitions.inclusion _ DP_R_xml_0_scc_38_large_large_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_38_large_large_large_in_le :
      Relation_Definitions.inclusion _ DP_R_xml_0_scc_38_large_large_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_38_large_large_large  := 
       WF_DP_R_xml_0_scc_38_large_large_large.wf.
     
     
     Lemma wf :
      well_founded WF_DP_R_xml_0_scc_38_large.DP_R_xml_0_scc_38_large_large.
     Proof.
       intros x.
       apply (well_founded_ind wf_lt).
       clear x.
       intros x.
       pattern x.
       apply (@Acc_ind _ DP_R_xml_0_scc_38_large_large_large).
       clear x.
       intros x _ IHx IHx'.
       constructor.
       intros y H.
       
       destruct H;
        (apply IHx';apply DP_R_xml_0_scc_38_large_large_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_38_large_large_large_in_le;
           econstructor eassumption])).
       apply wf_DP_R_xml_0_scc_38_large_large_large.
     Qed.
    End WF_DP_R_xml_0_scc_38_large_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_active (x8:Z) := 0.
    
    Definition P_id_add (x8:Z) (x9:Z) := 1.
    
    Definition P_id_recip (x8:Z) := 0.
    
    Definition P_id_mark (x8:Z) := 0.
    
    Definition P_id_first (x8:Z) (x9:Z) := 1.
    
    Definition P_id_s (x8:Z) := 0.
    
    Definition P_id_terms (x8:Z) := 1.
    
    Definition P_id_dbl (x8:Z) := 1.
    
    Definition P_id_sqr (x8:Z) := 1.
    
    Definition P_id_cons (x8:Z) (x9:Z) := 0.
    
    Definition P_id_nil  := 0.
    
    Definition P_id_0  := 0.
    
    Lemma P_id_active_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
    Proof.
      intros x9 x8.
      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_add_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_recip_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
    Proof.
      intros x9 x8.
      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_mark_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
    Proof.
      intros x9 x8.
      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_first_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_s_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
    Proof.
      intros x9 x8.
      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_terms_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
    Proof.
      intros x9 x8.
      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_dbl_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
    Proof.
      intros x9 x8.
      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_sqr_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
    Proof.
      intros x9 x8.
      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_cons_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_add_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_first_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_cons_bounded :
     forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_nil_bounded : 0 <= P_id_nil .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_0_bounded : 0 <= P_id_0 .
    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_active P_id_add P_id_recip P_id_mark P_id_first 
       P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil P_id_0.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                     P_id_active (measure x8)
                    | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                     P_id_add (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                     P_id_recip (measure x8)
                    | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                     P_id_mark (measure x8)
                    | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                     P_id_first (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                     P_id_s (measure x8)
                    | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                     P_id_terms (measure x8)
                    | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                     P_id_dbl (measure x8)
                    | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                     P_id_sqr (measure x8)
                    | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                     P_id_cons (measure x9) (measure x8)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                    | _ => 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_active_monotonic;assumption.
      intros ;apply P_id_add_monotonic;assumption.
      intros ;apply P_id_recip_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_first_monotonic;assumption.
      intros ;apply P_id_s_monotonic;assumption.
      intros ;apply P_id_terms_monotonic;assumption.
      intros ;apply P_id_dbl_monotonic;assumption.
      intros ;apply P_id_sqr_monotonic;assumption.
      intros ;apply P_id_cons_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_add_bounded;assumption.
      intros ;apply P_id_recip_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_first_bounded;assumption.
      intros ;apply P_id_s_bounded;assumption.
      intros ;apply P_id_terms_bounded;assumption.
      intros ;apply P_id_dbl_bounded;assumption.
      intros ;apply P_id_sqr_bounded;assumption.
      intros ;apply P_id_cons_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_0_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_SQR (x8:Z) := 0.
    
    Definition P_id_DBL (x8:Z) := 0.
    
    Definition P_id_ACTIVE (x8:Z) := 1* x8.
    
    Definition P_id_S (x8:Z) := 0.
    
    Definition P_id_CONS (x8:Z) (x9:Z) := 0.
    
    Definition P_id_TERMS (x8:Z) := 0.
    
    Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
    
    Definition P_id_MARK (x8:Z) := 1.
    
    Definition P_id_ADD (x8:Z) (x9:Z) := 0.
    
    Definition P_id_RECIP (x8:Z) := 0.
    
    Lemma P_id_SQR_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
    Proof.
      intros x9 x8.
      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_DBL_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
    Proof.
      intros x9 x8.
      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_ACTIVE_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
    Proof.
      intros x9 x8.
      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_S_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
    Proof.
      intros x9 x8.
      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_CONS_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_TERMS_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
    Proof.
      intros x9 x8.
      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_FIRST_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_MARK_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
    Proof.
      intros x9 x8.
      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_ADD_monotonic :
     forall x8 x10 x9 x11, 
      (0 <= x11)/\ (x11 <= x10) ->
       (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
    Proof.
      intros x11 x10 x9 x8.
      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_RECIP_monotonic :
     forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
    Proof.
      intros x9 x8.
      intros [H_1 H_0].
      
      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_active P_id_add P_id_recip P_id_mark 
       P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
       P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
       P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                            P_id_SQR (measure x8)
                           | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                            P_id_DBL (measure x8)
                           | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                            P_id_ACTIVE (measure x8)
                           | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                            P_id_S (measure x8)
                           | (algebra.Alg.Term algebra.F.id_cons (x9::
                              x8::nil)) =>
                            P_id_CONS (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                            P_id_TERMS (measure x8)
                           | (algebra.Alg.Term algebra.F.id_first (x9::
                              x8::nil)) =>
                            P_id_FIRST (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                            P_id_MARK (measure x8)
                           | (algebra.Alg.Term algebra.F.id_add (x9::
                              x8::nil)) =>
                            P_id_ADD (measure x9) (measure x8)
                           | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                            P_id_RECIP (measure x8)
                           | _ => 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_active_monotonic;assumption.
      intros ;apply P_id_add_monotonic;assumption.
      intros ;apply P_id_recip_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_first_monotonic;assumption.
      intros ;apply P_id_s_monotonic;assumption.
      intros ;apply P_id_terms_monotonic;assumption.
      intros ;apply P_id_dbl_monotonic;assumption.
      intros ;apply P_id_sqr_monotonic;assumption.
      intros ;apply P_id_cons_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_add_bounded;assumption.
      intros ;apply P_id_recip_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_first_bounded;assumption.
      intros ;apply P_id_s_bounded;assumption.
      intros ;apply P_id_terms_bounded;assumption.
      intros ;apply P_id_dbl_bounded;assumption.
      intros ;apply P_id_sqr_bounded;assumption.
      intros ;apply P_id_cons_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_0_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_SQR_monotonic;assumption.
      intros ;apply P_id_DBL_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_S_monotonic;assumption.
      intros ;apply P_id_CONS_monotonic;assumption.
      intros ;apply P_id_TERMS_monotonic;assumption.
      intros ;apply P_id_FIRST_monotonic;assumption.
      intros ;apply P_id_MARK_monotonic;assumption.
      intros ;apply P_id_ADD_monotonic;assumption.
      intros ;apply P_id_RECIP_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_38_large_strict_in_lt :
     Relation_Definitions.inclusion _ DP_R_xml_0_scc_38_large_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_38_large_large_in_le :
     Relation_Definitions.inclusion _ DP_R_xml_0_scc_38_large_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_38_large_large  := 
      WF_DP_R_xml_0_scc_38_large_large.wf.
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_38.DP_R_xml_0_scc_38_large.
    Proof.
      intros x.
      apply (well_founded_ind wf_lt).
      clear x.
      intros x.
      pattern x.
      apply (@Acc_ind _ DP_R_xml_0_scc_38_large_large).
      clear x.
      intros x _ IHx IHx'.
      constructor.
      intros y H.
      
      destruct H;
       (apply IHx';apply DP_R_xml_0_scc_38_large_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_38_large_large_in_le;econstructor eassumption])).
      apply wf_DP_R_xml_0_scc_38_large_large.
    Qed.
   End WF_DP_R_xml_0_scc_38_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_active (x8:Z) := 0.
   
   Definition P_id_add (x8:Z) (x9:Z) := 1.
   
   Definition P_id_recip (x8:Z) := 1.
   
   Definition P_id_mark (x8:Z) := 0.
   
   Definition P_id_first (x8:Z) (x9:Z) := 1.
   
   Definition P_id_s (x8:Z) := 0.
   
   Definition P_id_terms (x8:Z) := 1.
   
   Definition P_id_dbl (x8:Z) := 1.
   
   Definition P_id_sqr (x8:Z) := 1.
   
   Definition P_id_cons (x8:Z) (x9:Z) := 1.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_0  := 0.
   
   Lemma P_id_active_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_active x9 <= P_id_active x8.
   Proof.
     intros x9 x8.
     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_add_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_add x9 x11 <= P_id_add x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_recip_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_recip x9 <= P_id_recip x8.
   Proof.
     intros x9 x8.
     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_mark_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_mark x9 <= P_id_mark x8.
   Proof.
     intros x9 x8.
     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_first_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_first x9 x11 <= P_id_first x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_s_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_s x9 <= P_id_s x8.
   Proof.
     intros x9 x8.
     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_terms_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_terms x9 <= P_id_terms x8.
   Proof.
     intros x9 x8.
     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_dbl_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_dbl x9 <= P_id_dbl x8.
   Proof.
     intros x9 x8.
     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_sqr_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_sqr x9 <= P_id_sqr x8.
   Proof.
     intros x9 x8.
     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_cons_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_cons x9 x11 <= P_id_cons x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_active_bounded : forall x8, (0 <= x8) ->0 <= P_id_active x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_add_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_add x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_recip_bounded : forall x8, (0 <= x8) ->0 <= P_id_recip x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_mark_bounded : forall x8, (0 <= x8) ->0 <= P_id_mark x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_first_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_first x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_s_bounded : forall x8, (0 <= x8) ->0 <= P_id_s x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_terms_bounded : forall x8, (0 <= x8) ->0 <= P_id_terms x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_dbl_bounded : forall x8, (0 <= x8) ->0 <= P_id_dbl x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_sqr_bounded : forall x8, (0 <= x8) ->0 <= P_id_sqr x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_cons_bounded :
    forall x8 x9, (0 <= x8) ->(0 <= x9) ->0 <= P_id_cons x9 x8.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_nil_bounded : 0 <= P_id_nil .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_0_bounded : 0 <= P_id_0 .
   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_active P_id_add P_id_recip P_id_mark P_id_first 
      P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil P_id_0.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                    P_id_active (measure x8)
                   | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                    P_id_add (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                    P_id_recip (measure x8)
                   | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                    P_id_mark (measure x8)
                   | (algebra.Alg.Term algebra.F.id_first (x9::x8::nil)) =>
                    P_id_first (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                    P_id_s (measure x8)
                   | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                    P_id_terms (measure x8)
                   | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                    P_id_dbl (measure x8)
                   | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                    P_id_sqr (measure x8)
                   | (algebra.Alg.Term algebra.F.id_cons (x9::x8::nil)) =>
                    P_id_cons (measure x9) (measure x8)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 
                   | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_SQR (x8:Z) := 0.
   
   Definition P_id_DBL (x8:Z) := 0.
   
   Definition P_id_ACTIVE (x8:Z) := 1* x8.
   
   Definition P_id_S (x8:Z) := 0.
   
   Definition P_id_CONS (x8:Z) (x9:Z) := 0.
   
   Definition P_id_TERMS (x8:Z) := 0.
   
   Definition P_id_FIRST (x8:Z) (x9:Z) := 0.
   
   Definition P_id_MARK (x8:Z) := 1.
   
   Definition P_id_ADD (x8:Z) (x9:Z) := 0.
   
   Definition P_id_RECIP (x8:Z) := 0.
   
   Lemma P_id_SQR_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_SQR x9 <= P_id_SQR x8.
   Proof.
     intros x9 x8.
     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_DBL_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_DBL x9 <= P_id_DBL x8.
   Proof.
     intros x9 x8.
     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_ACTIVE_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_ACTIVE x9 <= P_id_ACTIVE x8.
   Proof.
     intros x9 x8.
     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_S_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_S x9 <= P_id_S x8.
   Proof.
     intros x9 x8.
     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_CONS_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_CONS x9 x11 <= P_id_CONS x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_TERMS_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_TERMS x9 <= P_id_TERMS x8.
   Proof.
     intros x9 x8.
     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_FIRST_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_FIRST x9 x11 <= P_id_FIRST x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_MARK_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_MARK x9 <= P_id_MARK x8.
   Proof.
     intros x9 x8.
     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_ADD_monotonic :
    forall x8 x10 x9 x11, 
     (0 <= x11)/\ (x11 <= x10) ->
      (0 <= x9)/\ (x9 <= x8) ->P_id_ADD x9 x11 <= P_id_ADD x8 x10.
   Proof.
     intros x11 x10 x9 x8.
     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_RECIP_monotonic :
    forall x8 x9, (0 <= x9)/\ (x9 <= x8) ->P_id_RECIP x9 <= P_id_RECIP x8.
   Proof.
     intros x9 x8.
     intros [H_1 H_0].
     
     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_active P_id_add P_id_recip P_id_mark 
      P_id_first P_id_s P_id_terms P_id_dbl P_id_sqr P_id_cons P_id_nil 
      P_id_0 P_id_SQR P_id_DBL P_id_ACTIVE P_id_S P_id_CONS P_id_TERMS 
      P_id_FIRST P_id_MARK P_id_ADD P_id_RECIP.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_sqr (x8::nil)) =>
                           P_id_SQR (measure x8)
                          | (algebra.Alg.Term algebra.F.id_dbl (x8::nil)) =>
                           P_id_DBL (measure x8)
                          | (algebra.Alg.Term algebra.F.id_active (x8::nil)) =>
                           P_id_ACTIVE (measure x8)
                          | (algebra.Alg.Term algebra.F.id_s (x8::nil)) =>
                           P_id_S (measure x8)
                          | (algebra.Alg.Term algebra.F.id_cons (x9::
                             x8::nil)) =>
                           P_id_CONS (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_terms (x8::nil)) =>
                           P_id_TERMS (measure x8)
                          | (algebra.Alg.Term algebra.F.id_first (x9::
                             x8::nil)) =>
                           P_id_FIRST (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_mark (x8::nil)) =>
                           P_id_MARK (measure x8)
                          | (algebra.Alg.Term algebra.F.id_add (x9::x8::nil)) =>
                           P_id_ADD (measure x9) (measure x8)
                          | (algebra.Alg.Term algebra.F.id_recip (x8::nil)) =>
                           P_id_RECIP (measure x8)
                          | _ => 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_active_monotonic;assumption.
     intros ;apply P_id_add_monotonic;assumption.
     intros ;apply P_id_recip_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_first_monotonic;assumption.
     intros ;apply P_id_s_monotonic;assumption.
     intros ;apply P_id_terms_monotonic;assumption.
     intros ;apply P_id_dbl_monotonic;assumption.
     intros ;apply P_id_sqr_monotonic;assumption.
     intros ;apply P_id_cons_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_add_bounded;assumption.
     intros ;apply P_id_recip_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_first_bounded;assumption.
     intros ;apply P_id_s_bounded;assumption.
     intros ;apply P_id_terms_bounded;assumption.
     intros ;apply P_id_dbl_bounded;assumption.
     intros ;apply P_id_sqr_bounded;assumption.
     intros ;apply P_id_cons_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_0_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_SQR_monotonic;assumption.
     intros ;apply P_id_DBL_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_S_monotonic;assumption.
     intros ;apply P_id_CONS_monotonic;assumption.
     intros ;apply P_id_TERMS_monotonic;assumption.
     intros ;apply P_id_FIRST_monotonic;assumption.
     intros ;apply P_id_MARK_monotonic;assumption.
     intros ;apply P_id_ADD_monotonic;assumption.
     intros ;apply P_id_RECIP_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_38_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_38_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_38_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_38_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_38_large  := WF_DP_R_xml_0_scc_38_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_38.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_38_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_38_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_38_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_38_large.
   Qed.
  End WF_DP_R_xml_0_scc_38.
  
  Definition wf_DP_R_xml_0_scc_38  := WF_DP_R_xml_0_scc_38.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_38 :
   forall x y, (DP_R_xml_0_scc_38 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_38).
    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_37;
         econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
       ((eapply acc_DP_R_xml_0_non_scc_36;
          econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
        ((eapply acc_DP_R_xml_0_non_scc_35;
           econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
         ((eapply acc_DP_R_xml_0_non_scc_34;
            econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
          ((eapply acc_DP_R_xml_0_non_scc_33;
             econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
           ((eapply acc_DP_R_xml_0_non_scc_32;
              econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
            ((eapply acc_DP_R_xml_0_non_scc_31;
               econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
             ((eapply acc_DP_R_xml_0_non_scc_30;
                econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
              ((eapply acc_DP_R_xml_0_non_scc_29;
                 econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
               ((eapply acc_DP_R_xml_0_non_scc_28;
                  econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                ((eapply acc_DP_R_xml_0_non_scc_27;
                   econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                 ((eapply acc_DP_R_xml_0_non_scc_26;
                    econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                  ((eapply acc_DP_R_xml_0_non_scc_25;
                     econstructor 
                     (eassumption)||(algebra.Alg_ext.star_refl' ))||
                   ((eapply acc_DP_R_xml_0_non_scc_24;
                      econstructor 
                      (eassumption)||(algebra.Alg_ext.star_refl' ))||
                    ((eapply acc_DP_R_xml_0_non_scc_23;
                       econstructor 
                       (eassumption)||(algebra.Alg_ext.star_refl' ))||
                     ((eapply acc_DP_R_xml_0_non_scc_22;
                        econstructor 
                        (eassumption)||(algebra.Alg_ext.star_refl' ))||
                      ((eapply acc_DP_R_xml_0_non_scc_21;
                         econstructor 
                         (eassumption)||(algebra.Alg_ext.star_refl' ))||
                       ((eapply acc_DP_R_xml_0_non_scc_20;
                          econstructor 
                          (eassumption)||(algebra.Alg_ext.star_refl' ))||
                        ((eapply acc_DP_R_xml_0_non_scc_19;
                           econstructor 
                           (eassumption)||(algebra.Alg_ext.star_refl' ))||
                         ((eapply acc_DP_R_xml_0_non_scc_18;
                            econstructor 
                            (eassumption)||(algebra.Alg_ext.star_refl' ))||
                          ((eapply acc_DP_R_xml_0_non_scc_17;
                             econstructor 
                             (eassumption)||(algebra.Alg_ext.star_refl' ))||
                           ((eapply acc_DP_R_xml_0_non_scc_16;
                              econstructor 
                              (eassumption)||(algebra.Alg_ext.star_refl' ))||
                            ((eapply acc_DP_R_xml_0_non_scc_15;
                               econstructor 
                               (eassumption)||(algebra.Alg_ext.star_refl' ))||
                             ((eapply acc_DP_R_xml_0_non_scc_14;
                                econstructor 
                                (eassumption)||(algebra.Alg_ext.star_refl' ))||
                              ((eapply acc_DP_R_xml_0_non_scc_13;
                                 econstructor 
                                 (eassumption)||(algebra.Alg_ext.star_refl' ))||
                               ((eapply acc_DP_R_xml_0_non_scc_12;
                                  econstructor 
                                  (eassumption)||
                                  (algebra.Alg_ext.star_refl' ))||
                                ((eapply acc_DP_R_xml_0_non_scc_11;
                                   econstructor 
                                   (eassumption)||
                                   (algebra.Alg_ext.star_refl' ))||
                                 ((eapply acc_DP_R_xml_0_non_scc_10;
                                    econstructor 
                                    (eassumption)||
                                    (algebra.Alg_ext.star_refl' ))||
                                  ((eapply acc_DP_R_xml_0_non_scc_9;
                                     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_38.
  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_37;
       econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
     ((eapply acc_DP_R_xml_0_non_scc_36;
        econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
      ((eapply acc_DP_R_xml_0_non_scc_35;
         econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
       ((eapply acc_DP_R_xml_0_non_scc_34;
          econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
        ((eapply acc_DP_R_xml_0_non_scc_33;
           econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
         ((eapply acc_DP_R_xml_0_non_scc_32;
            econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
          ((eapply acc_DP_R_xml_0_non_scc_31;
             econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
           ((eapply acc_DP_R_xml_0_non_scc_30;
              econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
            ((eapply acc_DP_R_xml_0_non_scc_29;
               econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
             ((eapply acc_DP_R_xml_0_non_scc_28;
                econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
              ((eapply acc_DP_R_xml_0_non_scc_27;
                 econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
               ((eapply acc_DP_R_xml_0_non_scc_26;
                  econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                ((eapply acc_DP_R_xml_0_non_scc_25;
                   econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                 ((eapply acc_DP_R_xml_0_non_scc_24;
                    econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                  ((eapply acc_DP_R_xml_0_non_scc_23;
                     econstructor 
                     (eassumption)||(algebra.Alg_ext.star_refl' ))||
                   ((eapply acc_DP_R_xml_0_non_scc_22;
                      econstructor 
                      (eassumption)||(algebra.Alg_ext.star_refl' ))||
                    ((eapply acc_DP_R_xml_0_non_scc_21;
                       econstructor 
                       (eassumption)||(algebra.Alg_ext.star_refl' ))||
                     ((eapply acc_DP_R_xml_0_non_scc_20;
                        econstructor 
                        (eassumption)||(algebra.Alg_ext.star_refl' ))||
                      ((eapply acc_DP_R_xml_0_non_scc_19;
                         econstructor 
                         (eassumption)||(algebra.Alg_ext.star_refl' ))||
                       ((eapply acc_DP_R_xml_0_non_scc_18;
                          econstructor 
                          (eassumption)||(algebra.Alg_ext.star_refl' ))||
                        ((eapply acc_DP_R_xml_0_non_scc_17;
                           econstructor 
                           (eassumption)||(algebra.Alg_ext.star_refl' ))||
                         ((eapply acc_DP_R_xml_0_non_scc_16;
                            econstructor 
                            (eassumption)||(algebra.Alg_ext.star_refl' ))||
                          ((eapply acc_DP_R_xml_0_non_scc_15;
                             econstructor 
                             (eassumption)||(algebra.Alg_ext.star_refl' ))||
                           ((eapply acc_DP_R_xml_0_non_scc_14;
                              econstructor 
                              (eassumption)||(algebra.Alg_ext.star_refl' ))||
                            ((eapply acc_DP_R_xml_0_non_scc_13;
                               econstructor 
                               (eassumption)||(algebra.Alg_ext.star_refl' ))||
                             ((eapply acc_DP_R_xml_0_non_scc_12;
                                econstructor 
                                (eassumption)||(algebra.Alg_ext.star_refl' ))||
                              ((eapply acc_DP_R_xml_0_non_scc_11;
                                 econstructor 
                                 (eassumption)||(algebra.Alg_ext.star_refl' ))||
                               ((eapply acc_DP_R_xml_0_non_scc_10;
                                  econstructor 
                                  (eassumption)||
                                  (algebra.Alg_ext.star_refl' ))||
                                ((eapply acc_DP_R_xml_0_non_scc_9;
                                   econstructor 
                                   (eassumption)||
                                   (algebra.Alg_ext.star_refl' ))||
                                 ((eapply acc_DP_R_xml_0_non_scc_8;
                                    econstructor 
                                    (eassumption)||
                                    (algebra.Alg_ext.star_refl' ))||
                                  ((eapply acc_DP_R_xml_0_non_scc_7;
                                     econstructor 
                                     (eassumption)||
                                     (algebra.Alg_ext.star_refl' ))||
                                   ((eapply acc_DP_R_xml_0_non_scc_6;
                                      econstructor 
                                      (eassumption)||
                                      (algebra.Alg_ext.star_refl' ))||
                                    ((eapply acc_DP_R_xml_0_non_scc_5;
                                       econstructor 
                                       (eassumption)||
                                       (algebra.Alg_ext.star_refl' ))||
                                     ((eapply acc_DP_R_xml_0_non_scc_4;
                                        econstructor 
                                        (eassumption)||
                                        (algebra.Alg_ext.star_refl' ))||
                                      ((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_38;
                                             econstructor 
                                             (eassumption)||
                                             (algebra.Alg_ext.star_refl' ))||
                                           ((eapply acc_DP_R_xml_0_scc_37;
                                              econstructor 
                                              (eassumption)||
                                              (algebra.Alg_ext.star_refl' ))||
                                            ((eapply acc_DP_R_xml_0_scc_36;
                                               econstructor 
                                               (eassumption)||
                                               (algebra.Alg_ext.star_refl' ))||
                                             ((eapply acc_DP_R_xml_0_scc_35;
                                                econstructor 
                                                (eassumption)||
                                                (algebra.Alg_ext.star_refl' ))||
                                              ((eapply acc_DP_R_xml_0_scc_34;
                                                 econstructor 
                                                 (eassumption)||
                                                 (algebra.Alg_ext.star_refl' ))||
                                               ((eapply acc_DP_R_xml_0_scc_33;
                                                  econstructor 
                                                  (eassumption)||
                                                  (algebra.Alg_ext.star_refl' ))||
                                                ((eapply acc_DP_R_xml_0_scc_32;
                                                   
                                                   econstructor 
                                                   (eassumption)||
                                                   (algebra.Alg_ext.star_refl' ))||
                                                 ((eapply acc_DP_R_xml_0_scc_31;
                                                    
                                                    econstructor 
                                                    (eassumption)||
                                                    (algebra.Alg_ext.star_refl' ))||
                                                  ((eapply acc_DP_R_xml_0_scc_30;
                                                     
                                                     econstructor 
                                                     (eassumption)||
                                                     (algebra.Alg_ext.star_refl' ))||
                                                   ((eapply acc_DP_R_xml_0_scc_29;
                                                      
                                                      econstructor 
                                                      (eassumption)||
                                                      (algebra.Alg_ext.star_refl' ))||
                                                    ((eapply acc_DP_R_xml_0_scc_28;
                                                       
                                                       econstructor 
                                                       (eassumption)||
                                                       (algebra.Alg_ext.star_refl' ))||
                                                     ((eapply acc_DP_R_xml_0_scc_27;
                                                        
                                                        econstructor 
                                                        (eassumption)||
                                                        (algebra.Alg_ext.star_refl' ))||
                                                      ((eapply acc_DP_R_xml_0_scc_26;
                                                         
                                                         econstructor 
                                                         (eassumption)||
                                                         (algebra.Alg_ext.star_refl' ))||
                                                       ((eapply acc_DP_R_xml_0_scc_25;
                                                          
                                                          econstructor 
                                                          (eassumption)||
                                                          (algebra.Alg_ext.star_refl' ))||
                                                        ((eapply acc_DP_R_xml_0_scc_24;
                                                           
                                                           econstructor 
                                                           (eassumption)||
                                                           (algebra.Alg_ext.star_refl' ))||
                                                         ((eapply acc_DP_R_xml_0_scc_23;
                                                            
                                                            econstructor 
                                                            (eassumption)||
                                                            (algebra.Alg_ext.star_refl' ))||
                                                          ((eapply acc_DP_R_xml_0_scc_22;
                                                             
                                                             econstructor 
                                                             (eassumption)||
                                                             (algebra.Alg_ext.star_refl' ))||
                                                           ((eapply acc_DP_R_xml_0_scc_21;
                                                              
                                                              econstructor 
                                                              (eassumption)||
                                                              (algebra.Alg_ext.star_refl' ))||
                                                            ((eapply 
                                                              acc_DP_R_xml_0_scc_20;
                                                               
                                                               econstructor 
                                                               (eassumption)||
                                                               (algebra.Alg_ext.star_refl' ))||
                                                             ((eapply 
                                                               acc_DP_R_xml_0_scc_19;
                                                                
                                                                econstructor 
                                                                (eassumption)||
                                                                (algebra.Alg_ext.star_refl' ))||
                                                              ((eapply 
                                                                acc_DP_R_xml_0_scc_18;
                                                                 
                                                                 econstructor 
                                                                 (eassumption)||
                                                                 (algebra.Alg_ext.star_refl' ))||
                                                               ((eapply 
                                                                 acc_DP_R_xml_0_scc_17;
                                                                  
                                                                  econstructor 
                                                                  (eassumption)||
                                                                  (algebra.Alg_ext.star_refl' ))||
                                                                ((eapply 
                                                                  acc_DP_R_xml_0_scc_16;
                                                                   
                                                                   econstructor 
                                                                   (eassumption)||
                                                                   (algebra.Alg_ext.star_refl' ))||
                                                                 ((eapply 
                                                                   acc_DP_R_xml_0_scc_15;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                  ((eapply 
                                                                    acc_DP_R_xml_0_scc_14;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                   ((
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_13;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_12;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_11;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_10;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_9;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_8;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_7;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_6;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_5;
                                                                    
                                                                    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___TRCSR___Ex26_Luc03b_iGM.trs/a3pat.v" ***
*** End: ***
 *)