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_U71 *)
    | id_U71 : symb
     (* id_isList *)
    | id_isList : symb
     (* id_i *)
    | id_i : symb
     (* id_U11 *)
    | id_U11 : symb
     (* id_isQid *)
    | id_isQid : symb
     (* id_isNeList *)
    | id_isNeList : symb
     (* id_ok *)
    | id_ok : symb
     (* id_mark *)
    | id_mark : symb
     (* id_isPal *)
    | id_isPal : symb
     (* id_U41 *)
    | id_U41 : symb
     (* id_u *)
    | id_u : symb
     (* id_U21 *)
    | id_U21 : symb
     (* id_a *)
    | id_a : symb
     (* id_U52 *)
    | id_U52 : symb
     (* id___ *)
    | id___ : symb
     (* id_U72 *)
    | id_U72 : symb
     (* id_U31 *)
    | id_U31 : symb
     (* id_o *)
    | id_o : symb
     (* id_tt *)
    | id_tt : symb
     (* id_isNePal *)
    | id_isNePal : symb
     (* id_U51 *)
    | id_U51 : symb
     (* id_top *)
    | id_top : symb
     (* id_nil *)
    | id_nil : symb
     (* id_U81 *)
    | id_U81 : symb
     (* id_U42 *)
    | id_U42 : symb
     (* id_proper *)
    | id_proper : symb
     (* id_U22 *)
    | id_U22 : symb
     (* id_e *)
    | id_e : symb
     (* id_U61 *)
    | id_U61 : symb
  .
  
  
  Definition symb_eq_bool (f1 f2:symb) : bool := 
    match f1,f2 with
      | id_active,id_active => true
      | id_U71,id_U71 => true
      | id_isList,id_isList => true
      | id_i,id_i => true
      | id_U11,id_U11 => true
      | id_isQid,id_isQid => true
      | id_isNeList,id_isNeList => true
      | id_ok,id_ok => true
      | id_mark,id_mark => true
      | id_isPal,id_isPal => true
      | id_U41,id_U41 => true
      | id_u,id_u => true
      | id_U21,id_U21 => true
      | id_a,id_a => true
      | id_U52,id_U52 => true
      | id___,id___ => true
      | id_U72,id_U72 => true
      | id_U31,id_U31 => true
      | id_o,id_o => true
      | id_tt,id_tt => true
      | id_isNePal,id_isNePal => true
      | id_U51,id_U51 => true
      | id_top,id_top => true
      | id_nil,id_nil => true
      | id_U81,id_U81 => true
      | id_U42,id_U42 => true
      | id_proper,id_proper => true
      | id_U22,id_U22 => true
      | id_e,id_e => true
      | id_U61,id_U61 => 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_U71,id_U71 => refl_equal _
             | id_isList,id_isList => refl_equal _
             | id_i,id_i => refl_equal _
             | id_U11,id_U11 => refl_equal _
             | id_isQid,id_isQid => refl_equal _
             | id_isNeList,id_isNeList => refl_equal _
             | id_ok,id_ok => refl_equal _
             | id_mark,id_mark => refl_equal _
             | id_isPal,id_isPal => refl_equal _
             | id_U41,id_U41 => refl_equal _
             | id_u,id_u => refl_equal _
             | id_U21,id_U21 => refl_equal _
             | id_a,id_a => refl_equal _
             | id_U52,id_U52 => refl_equal _
             | id___,id___ => refl_equal _
             | id_U72,id_U72 => refl_equal _
             | id_U31,id_U31 => refl_equal _
             | id_o,id_o => refl_equal _
             | id_tt,id_tt => refl_equal _
             | id_isNePal,id_isNePal => refl_equal _
             | id_U51,id_U51 => refl_equal _
             | id_top,id_top => refl_equal _
             | id_nil,id_nil => refl_equal _
             | id_U81,id_U81 => refl_equal _
             | id_U42,id_U42 => refl_equal _
             | id_proper,id_proper => refl_equal _
             | id_U22,id_U22 => refl_equal _
             | id_e,id_e => refl_equal _
             | id_U61,id_U61 => refl_equal _
             | _,_ => _
             end;intros abs;discriminate.
  Defined.
  
  
  Definition arity (f:symb) := 
    match f with
      | id_active => term.Free 1
      | id_U71 => term.Free 2
      | id_isList => term.Free 1
      | id_i => term.Free 0
      | id_U11 => term.Free 1
      | id_isQid => term.Free 1
      | id_isNeList => term.Free 1
      | id_ok => term.Free 1
      | id_mark => term.Free 1
      | id_isPal => term.Free 1
      | id_U41 => term.Free 2
      | id_u => term.Free 0
      | id_U21 => term.Free 2
      | id_a => term.Free 0
      | id_U52 => term.Free 1
      | id___ => term.Free 2
      | id_U72 => term.Free 1
      | id_U31 => term.Free 1
      | id_o => term.Free 0
      | id_tt => term.Free 0
      | id_isNePal => term.Free 1
      | id_U51 => term.Free 2
      | id_top => term.Free 1
      | id_nil => term.Free 0
      | id_U81 => term.Free 1
      | id_U42 => term.Free 1
      | id_proper => term.Free 1
      | id_U22 => term.Free 1
      | id_e => term.Free 0
      | id_U61 => term.Free 1
      end.
  
  
  Definition symb_order (f1 f2:symb) : bool := 
    match f1,f2 with
      | id_active,id_active => true
      | id_active,id_U71 => false
      | id_active,id_isList => false
      | id_active,id_i => false
      | id_active,id_U11 => false
      | id_active,id_isQid => false
      | id_active,id_isNeList => false
      | id_active,id_ok => false
      | id_active,id_mark => false
      | id_active,id_isPal => false
      | id_active,id_U41 => false
      | id_active,id_u => false
      | id_active,id_U21 => false
      | id_active,id_a => false
      | id_active,id_U52 => false
      | id_active,id___ => false
      | id_active,id_U72 => false
      | id_active,id_U31 => false
      | id_active,id_o => false
      | id_active,id_tt => false
      | id_active,id_isNePal => false
      | id_active,id_U51 => false
      | id_active,id_top => false
      | id_active,id_nil => false
      | id_active,id_U81 => false
      | id_active,id_U42 => false
      | id_active,id_proper => false
      | id_active,id_U22 => false
      | id_active,id_e => false
      | id_active,id_U61 => false
      | id_U71,id_active => true
      | id_U71,id_U71 => true
      | id_U71,id_isList => false
      | id_U71,id_i => false
      | id_U71,id_U11 => false
      | id_U71,id_isQid => false
      | id_U71,id_isNeList => false
      | id_U71,id_ok => false
      | id_U71,id_mark => false
      | id_U71,id_isPal => false
      | id_U71,id_U41 => false
      | id_U71,id_u => false
      | id_U71,id_U21 => false
      | id_U71,id_a => false
      | id_U71,id_U52 => false
      | id_U71,id___ => false
      | id_U71,id_U72 => false
      | id_U71,id_U31 => false
      | id_U71,id_o => false
      | id_U71,id_tt => false
      | id_U71,id_isNePal => false
      | id_U71,id_U51 => false
      | id_U71,id_top => false
      | id_U71,id_nil => false
      | id_U71,id_U81 => false
      | id_U71,id_U42 => false
      | id_U71,id_proper => false
      | id_U71,id_U22 => false
      | id_U71,id_e => false
      | id_U71,id_U61 => false
      | id_isList,id_active => true
      | id_isList,id_U71 => true
      | id_isList,id_isList => true
      | id_isList,id_i => false
      | id_isList,id_U11 => false
      | id_isList,id_isQid => false
      | id_isList,id_isNeList => false
      | id_isList,id_ok => false
      | id_isList,id_mark => false
      | id_isList,id_isPal => false
      | id_isList,id_U41 => false
      | id_isList,id_u => false
      | id_isList,id_U21 => false
      | id_isList,id_a => false
      | id_isList,id_U52 => false
      | id_isList,id___ => false
      | id_isList,id_U72 => false
      | id_isList,id_U31 => false
      | id_isList,id_o => false
      | id_isList,id_tt => false
      | id_isList,id_isNePal => false
      | id_isList,id_U51 => false
      | id_isList,id_top => false
      | id_isList,id_nil => false
      | id_isList,id_U81 => false
      | id_isList,id_U42 => false
      | id_isList,id_proper => false
      | id_isList,id_U22 => false
      | id_isList,id_e => false
      | id_isList,id_U61 => false
      | id_i,id_active => true
      | id_i,id_U71 => true
      | id_i,id_isList => true
      | id_i,id_i => true
      | id_i,id_U11 => false
      | id_i,id_isQid => false
      | id_i,id_isNeList => false
      | id_i,id_ok => false
      | id_i,id_mark => false
      | id_i,id_isPal => false
      | id_i,id_U41 => false
      | id_i,id_u => false
      | id_i,id_U21 => false
      | id_i,id_a => false
      | id_i,id_U52 => false
      | id_i,id___ => false
      | id_i,id_U72 => false
      | id_i,id_U31 => false
      | id_i,id_o => false
      | id_i,id_tt => false
      | id_i,id_isNePal => false
      | id_i,id_U51 => false
      | id_i,id_top => false
      | id_i,id_nil => false
      | id_i,id_U81 => false
      | id_i,id_U42 => false
      | id_i,id_proper => false
      | id_i,id_U22 => false
      | id_i,id_e => false
      | id_i,id_U61 => false
      | id_U11,id_active => true
      | id_U11,id_U71 => true
      | id_U11,id_isList => true
      | id_U11,id_i => true
      | id_U11,id_U11 => true
      | id_U11,id_isQid => false
      | id_U11,id_isNeList => false
      | id_U11,id_ok => false
      | id_U11,id_mark => false
      | id_U11,id_isPal => false
      | id_U11,id_U41 => false
      | id_U11,id_u => false
      | id_U11,id_U21 => false
      | id_U11,id_a => false
      | id_U11,id_U52 => false
      | id_U11,id___ => false
      | id_U11,id_U72 => false
      | id_U11,id_U31 => false
      | id_U11,id_o => false
      | id_U11,id_tt => false
      | id_U11,id_isNePal => false
      | id_U11,id_U51 => false
      | id_U11,id_top => false
      | id_U11,id_nil => false
      | id_U11,id_U81 => false
      | id_U11,id_U42 => false
      | id_U11,id_proper => false
      | id_U11,id_U22 => false
      | id_U11,id_e => false
      | id_U11,id_U61 => false
      | id_isQid,id_active => true
      | id_isQid,id_U71 => true
      | id_isQid,id_isList => true
      | id_isQid,id_i => true
      | id_isQid,id_U11 => true
      | id_isQid,id_isQid => true
      | id_isQid,id_isNeList => false
      | id_isQid,id_ok => false
      | id_isQid,id_mark => false
      | id_isQid,id_isPal => false
      | id_isQid,id_U41 => false
      | id_isQid,id_u => false
      | id_isQid,id_U21 => false
      | id_isQid,id_a => false
      | id_isQid,id_U52 => false
      | id_isQid,id___ => false
      | id_isQid,id_U72 => false
      | id_isQid,id_U31 => false
      | id_isQid,id_o => false
      | id_isQid,id_tt => false
      | id_isQid,id_isNePal => false
      | id_isQid,id_U51 => false
      | id_isQid,id_top => false
      | id_isQid,id_nil => false
      | id_isQid,id_U81 => false
      | id_isQid,id_U42 => false
      | id_isQid,id_proper => false
      | id_isQid,id_U22 => false
      | id_isQid,id_e => false
      | id_isQid,id_U61 => false
      | id_isNeList,id_active => true
      | id_isNeList,id_U71 => true
      | id_isNeList,id_isList => true
      | id_isNeList,id_i => true
      | id_isNeList,id_U11 => true
      | id_isNeList,id_isQid => true
      | id_isNeList,id_isNeList => true
      | id_isNeList,id_ok => false
      | id_isNeList,id_mark => false
      | id_isNeList,id_isPal => false
      | id_isNeList,id_U41 => false
      | id_isNeList,id_u => false
      | id_isNeList,id_U21 => false
      | id_isNeList,id_a => false
      | id_isNeList,id_U52 => false
      | id_isNeList,id___ => false
      | id_isNeList,id_U72 => false
      | id_isNeList,id_U31 => false
      | id_isNeList,id_o => false
      | id_isNeList,id_tt => false
      | id_isNeList,id_isNePal => false
      | id_isNeList,id_U51 => false
      | id_isNeList,id_top => false
      | id_isNeList,id_nil => false
      | id_isNeList,id_U81 => false
      | id_isNeList,id_U42 => false
      | id_isNeList,id_proper => false
      | id_isNeList,id_U22 => false
      | id_isNeList,id_e => false
      | id_isNeList,id_U61 => false
      | id_ok,id_active => true
      | id_ok,id_U71 => true
      | id_ok,id_isList => true
      | id_ok,id_i => true
      | id_ok,id_U11 => true
      | id_ok,id_isQid => true
      | id_ok,id_isNeList => true
      | id_ok,id_ok => true
      | id_ok,id_mark => false
      | id_ok,id_isPal => false
      | id_ok,id_U41 => false
      | id_ok,id_u => false
      | id_ok,id_U21 => false
      | id_ok,id_a => false
      | id_ok,id_U52 => false
      | id_ok,id___ => false
      | id_ok,id_U72 => false
      | id_ok,id_U31 => false
      | id_ok,id_o => false
      | id_ok,id_tt => false
      | id_ok,id_isNePal => false
      | id_ok,id_U51 => false
      | id_ok,id_top => false
      | id_ok,id_nil => false
      | id_ok,id_U81 => false
      | id_ok,id_U42 => false
      | id_ok,id_proper => false
      | id_ok,id_U22 => false
      | id_ok,id_e => false
      | id_ok,id_U61 => false
      | id_mark,id_active => true
      | id_mark,id_U71 => true
      | id_mark,id_isList => true
      | id_mark,id_i => true
      | id_mark,id_U11 => true
      | id_mark,id_isQid => true
      | id_mark,id_isNeList => true
      | id_mark,id_ok => true
      | id_mark,id_mark => true
      | id_mark,id_isPal => false
      | id_mark,id_U41 => false
      | id_mark,id_u => false
      | id_mark,id_U21 => false
      | id_mark,id_a => false
      | id_mark,id_U52 => false
      | id_mark,id___ => false
      | id_mark,id_U72 => false
      | id_mark,id_U31 => false
      | id_mark,id_o => false
      | id_mark,id_tt => false
      | id_mark,id_isNePal => false
      | id_mark,id_U51 => false
      | id_mark,id_top => false
      | id_mark,id_nil => false
      | id_mark,id_U81 => false
      | id_mark,id_U42 => false
      | id_mark,id_proper => false
      | id_mark,id_U22 => false
      | id_mark,id_e => false
      | id_mark,id_U61 => false
      | id_isPal,id_active => true
      | id_isPal,id_U71 => true
      | id_isPal,id_isList => true
      | id_isPal,id_i => true
      | id_isPal,id_U11 => true
      | id_isPal,id_isQid => true
      | id_isPal,id_isNeList => true
      | id_isPal,id_ok => true
      | id_isPal,id_mark => true
      | id_isPal,id_isPal => true
      | id_isPal,id_U41 => false
      | id_isPal,id_u => false
      | id_isPal,id_U21 => false
      | id_isPal,id_a => false
      | id_isPal,id_U52 => false
      | id_isPal,id___ => false
      | id_isPal,id_U72 => false
      | id_isPal,id_U31 => false
      | id_isPal,id_o => false
      | id_isPal,id_tt => false
      | id_isPal,id_isNePal => false
      | id_isPal,id_U51 => false
      | id_isPal,id_top => false
      | id_isPal,id_nil => false
      | id_isPal,id_U81 => false
      | id_isPal,id_U42 => false
      | id_isPal,id_proper => false
      | id_isPal,id_U22 => false
      | id_isPal,id_e => false
      | id_isPal,id_U61 => false
      | id_U41,id_active => true
      | id_U41,id_U71 => true
      | id_U41,id_isList => true
      | id_U41,id_i => true
      | id_U41,id_U11 => true
      | id_U41,id_isQid => true
      | id_U41,id_isNeList => true
      | id_U41,id_ok => true
      | id_U41,id_mark => true
      | id_U41,id_isPal => true
      | id_U41,id_U41 => true
      | id_U41,id_u => false
      | id_U41,id_U21 => false
      | id_U41,id_a => false
      | id_U41,id_U52 => false
      | id_U41,id___ => false
      | id_U41,id_U72 => false
      | id_U41,id_U31 => false
      | id_U41,id_o => false
      | id_U41,id_tt => false
      | id_U41,id_isNePal => false
      | id_U41,id_U51 => false
      | id_U41,id_top => false
      | id_U41,id_nil => false
      | id_U41,id_U81 => false
      | id_U41,id_U42 => false
      | id_U41,id_proper => false
      | id_U41,id_U22 => false
      | id_U41,id_e => false
      | id_U41,id_U61 => false
      | id_u,id_active => true
      | id_u,id_U71 => true
      | id_u,id_isList => true
      | id_u,id_i => true
      | id_u,id_U11 => true
      | id_u,id_isQid => true
      | id_u,id_isNeList => true
      | id_u,id_ok => true
      | id_u,id_mark => true
      | id_u,id_isPal => true
      | id_u,id_U41 => true
      | id_u,id_u => true
      | id_u,id_U21 => false
      | id_u,id_a => false
      | id_u,id_U52 => false
      | id_u,id___ => false
      | id_u,id_U72 => false
      | id_u,id_U31 => false
      | id_u,id_o => false
      | id_u,id_tt => false
      | id_u,id_isNePal => false
      | id_u,id_U51 => false
      | id_u,id_top => false
      | id_u,id_nil => false
      | id_u,id_U81 => false
      | id_u,id_U42 => false
      | id_u,id_proper => false
      | id_u,id_U22 => false
      | id_u,id_e => false
      | id_u,id_U61 => false
      | id_U21,id_active => true
      | id_U21,id_U71 => true
      | id_U21,id_isList => true
      | id_U21,id_i => true
      | id_U21,id_U11 => true
      | id_U21,id_isQid => true
      | id_U21,id_isNeList => true
      | id_U21,id_ok => true
      | id_U21,id_mark => true
      | id_U21,id_isPal => true
      | id_U21,id_U41 => true
      | id_U21,id_u => true
      | id_U21,id_U21 => true
      | id_U21,id_a => false
      | id_U21,id_U52 => false
      | id_U21,id___ => false
      | id_U21,id_U72 => false
      | id_U21,id_U31 => false
      | id_U21,id_o => false
      | id_U21,id_tt => false
      | id_U21,id_isNePal => false
      | id_U21,id_U51 => false
      | id_U21,id_top => false
      | id_U21,id_nil => false
      | id_U21,id_U81 => false
      | id_U21,id_U42 => false
      | id_U21,id_proper => false
      | id_U21,id_U22 => false
      | id_U21,id_e => false
      | id_U21,id_U61 => false
      | id_a,id_active => true
      | id_a,id_U71 => true
      | id_a,id_isList => true
      | id_a,id_i => true
      | id_a,id_U11 => true
      | id_a,id_isQid => true
      | id_a,id_isNeList => true
      | id_a,id_ok => true
      | id_a,id_mark => true
      | id_a,id_isPal => true
      | id_a,id_U41 => true
      | id_a,id_u => true
      | id_a,id_U21 => true
      | id_a,id_a => true
      | id_a,id_U52 => false
      | id_a,id___ => false
      | id_a,id_U72 => false
      | id_a,id_U31 => false
      | id_a,id_o => false
      | id_a,id_tt => false
      | id_a,id_isNePal => false
      | id_a,id_U51 => false
      | id_a,id_top => false
      | id_a,id_nil => false
      | id_a,id_U81 => false
      | id_a,id_U42 => false
      | id_a,id_proper => false
      | id_a,id_U22 => false
      | id_a,id_e => false
      | id_a,id_U61 => false
      | id_U52,id_active => true
      | id_U52,id_U71 => true
      | id_U52,id_isList => true
      | id_U52,id_i => true
      | id_U52,id_U11 => true
      | id_U52,id_isQid => true
      | id_U52,id_isNeList => true
      | id_U52,id_ok => true
      | id_U52,id_mark => true
      | id_U52,id_isPal => true
      | id_U52,id_U41 => true
      | id_U52,id_u => true
      | id_U52,id_U21 => true
      | id_U52,id_a => true
      | id_U52,id_U52 => true
      | id_U52,id___ => false
      | id_U52,id_U72 => false
      | id_U52,id_U31 => false
      | id_U52,id_o => false
      | id_U52,id_tt => false
      | id_U52,id_isNePal => false
      | id_U52,id_U51 => false
      | id_U52,id_top => false
      | id_U52,id_nil => false
      | id_U52,id_U81 => false
      | id_U52,id_U42 => false
      | id_U52,id_proper => false
      | id_U52,id_U22 => false
      | id_U52,id_e => false
      | id_U52,id_U61 => false
      | id___,id_active => true
      | id___,id_U71 => true
      | id___,id_isList => true
      | id___,id_i => true
      | id___,id_U11 => true
      | id___,id_isQid => true
      | id___,id_isNeList => true
      | id___,id_ok => true
      | id___,id_mark => true
      | id___,id_isPal => true
      | id___,id_U41 => true
      | id___,id_u => true
      | id___,id_U21 => true
      | id___,id_a => true
      | id___,id_U52 => true
      | id___,id___ => true
      | id___,id_U72 => false
      | id___,id_U31 => false
      | id___,id_o => false
      | id___,id_tt => false
      | id___,id_isNePal => false
      | id___,id_U51 => false
      | id___,id_top => false
      | id___,id_nil => false
      | id___,id_U81 => false
      | id___,id_U42 => false
      | id___,id_proper => false
      | id___,id_U22 => false
      | id___,id_e => false
      | id___,id_U61 => false
      | id_U72,id_active => true
      | id_U72,id_U71 => true
      | id_U72,id_isList => true
      | id_U72,id_i => true
      | id_U72,id_U11 => true
      | id_U72,id_isQid => true
      | id_U72,id_isNeList => true
      | id_U72,id_ok => true
      | id_U72,id_mark => true
      | id_U72,id_isPal => true
      | id_U72,id_U41 => true
      | id_U72,id_u => true
      | id_U72,id_U21 => true
      | id_U72,id_a => true
      | id_U72,id_U52 => true
      | id_U72,id___ => true
      | id_U72,id_U72 => true
      | id_U72,id_U31 => false
      | id_U72,id_o => false
      | id_U72,id_tt => false
      | id_U72,id_isNePal => false
      | id_U72,id_U51 => false
      | id_U72,id_top => false
      | id_U72,id_nil => false
      | id_U72,id_U81 => false
      | id_U72,id_U42 => false
      | id_U72,id_proper => false
      | id_U72,id_U22 => false
      | id_U72,id_e => false
      | id_U72,id_U61 => false
      | id_U31,id_active => true
      | id_U31,id_U71 => true
      | id_U31,id_isList => true
      | id_U31,id_i => true
      | id_U31,id_U11 => true
      | id_U31,id_isQid => true
      | id_U31,id_isNeList => true
      | id_U31,id_ok => true
      | id_U31,id_mark => true
      | id_U31,id_isPal => true
      | id_U31,id_U41 => true
      | id_U31,id_u => true
      | id_U31,id_U21 => true
      | id_U31,id_a => true
      | id_U31,id_U52 => true
      | id_U31,id___ => true
      | id_U31,id_U72 => true
      | id_U31,id_U31 => true
      | id_U31,id_o => false
      | id_U31,id_tt => false
      | id_U31,id_isNePal => false
      | id_U31,id_U51 => false
      | id_U31,id_top => false
      | id_U31,id_nil => false
      | id_U31,id_U81 => false
      | id_U31,id_U42 => false
      | id_U31,id_proper => false
      | id_U31,id_U22 => false
      | id_U31,id_e => false
      | id_U31,id_U61 => false
      | id_o,id_active => true
      | id_o,id_U71 => true
      | id_o,id_isList => true
      | id_o,id_i => true
      | id_o,id_U11 => true
      | id_o,id_isQid => true
      | id_o,id_isNeList => true
      | id_o,id_ok => true
      | id_o,id_mark => true
      | id_o,id_isPal => true
      | id_o,id_U41 => true
      | id_o,id_u => true
      | id_o,id_U21 => true
      | id_o,id_a => true
      | id_o,id_U52 => true
      | id_o,id___ => true
      | id_o,id_U72 => true
      | id_o,id_U31 => true
      | id_o,id_o => true
      | id_o,id_tt => false
      | id_o,id_isNePal => false
      | id_o,id_U51 => false
      | id_o,id_top => false
      | id_o,id_nil => false
      | id_o,id_U81 => false
      | id_o,id_U42 => false
      | id_o,id_proper => false
      | id_o,id_U22 => false
      | id_o,id_e => false
      | id_o,id_U61 => false
      | id_tt,id_active => true
      | id_tt,id_U71 => true
      | id_tt,id_isList => true
      | id_tt,id_i => true
      | id_tt,id_U11 => true
      | id_tt,id_isQid => true
      | id_tt,id_isNeList => true
      | id_tt,id_ok => true
      | id_tt,id_mark => true
      | id_tt,id_isPal => true
      | id_tt,id_U41 => true
      | id_tt,id_u => true
      | id_tt,id_U21 => true
      | id_tt,id_a => true
      | id_tt,id_U52 => true
      | id_tt,id___ => true
      | id_tt,id_U72 => true
      | id_tt,id_U31 => true
      | id_tt,id_o => true
      | id_tt,id_tt => true
      | id_tt,id_isNePal => false
      | id_tt,id_U51 => false
      | id_tt,id_top => false
      | id_tt,id_nil => false
      | id_tt,id_U81 => false
      | id_tt,id_U42 => false
      | id_tt,id_proper => false
      | id_tt,id_U22 => false
      | id_tt,id_e => false
      | id_tt,id_U61 => false
      | id_isNePal,id_active => true
      | id_isNePal,id_U71 => true
      | id_isNePal,id_isList => true
      | id_isNePal,id_i => true
      | id_isNePal,id_U11 => true
      | id_isNePal,id_isQid => true
      | id_isNePal,id_isNeList => true
      | id_isNePal,id_ok => true
      | id_isNePal,id_mark => true
      | id_isNePal,id_isPal => true
      | id_isNePal,id_U41 => true
      | id_isNePal,id_u => true
      | id_isNePal,id_U21 => true
      | id_isNePal,id_a => true
      | id_isNePal,id_U52 => true
      | id_isNePal,id___ => true
      | id_isNePal,id_U72 => true
      | id_isNePal,id_U31 => true
      | id_isNePal,id_o => true
      | id_isNePal,id_tt => true
      | id_isNePal,id_isNePal => true
      | id_isNePal,id_U51 => false
      | id_isNePal,id_top => false
      | id_isNePal,id_nil => false
      | id_isNePal,id_U81 => false
      | id_isNePal,id_U42 => false
      | id_isNePal,id_proper => false
      | id_isNePal,id_U22 => false
      | id_isNePal,id_e => false
      | id_isNePal,id_U61 => false
      | id_U51,id_active => true
      | id_U51,id_U71 => true
      | id_U51,id_isList => true
      | id_U51,id_i => true
      | id_U51,id_U11 => true
      | id_U51,id_isQid => true
      | id_U51,id_isNeList => true
      | id_U51,id_ok => true
      | id_U51,id_mark => true
      | id_U51,id_isPal => true
      | id_U51,id_U41 => true
      | id_U51,id_u => true
      | id_U51,id_U21 => true
      | id_U51,id_a => true
      | id_U51,id_U52 => true
      | id_U51,id___ => true
      | id_U51,id_U72 => true
      | id_U51,id_U31 => true
      | id_U51,id_o => true
      | id_U51,id_tt => true
      | id_U51,id_isNePal => true
      | id_U51,id_U51 => true
      | id_U51,id_top => false
      | id_U51,id_nil => false
      | id_U51,id_U81 => false
      | id_U51,id_U42 => false
      | id_U51,id_proper => false
      | id_U51,id_U22 => false
      | id_U51,id_e => false
      | id_U51,id_U61 => false
      | id_top,id_active => true
      | id_top,id_U71 => true
      | id_top,id_isList => true
      | id_top,id_i => true
      | id_top,id_U11 => true
      | id_top,id_isQid => true
      | id_top,id_isNeList => true
      | id_top,id_ok => true
      | id_top,id_mark => true
      | id_top,id_isPal => true
      | id_top,id_U41 => true
      | id_top,id_u => true
      | id_top,id_U21 => true
      | id_top,id_a => true
      | id_top,id_U52 => true
      | id_top,id___ => true
      | id_top,id_U72 => true
      | id_top,id_U31 => true
      | id_top,id_o => true
      | id_top,id_tt => true
      | id_top,id_isNePal => true
      | id_top,id_U51 => true
      | id_top,id_top => true
      | id_top,id_nil => false
      | id_top,id_U81 => false
      | id_top,id_U42 => false
      | id_top,id_proper => false
      | id_top,id_U22 => false
      | id_top,id_e => false
      | id_top,id_U61 => false
      | id_nil,id_active => true
      | id_nil,id_U71 => true
      | id_nil,id_isList => true
      | id_nil,id_i => true
      | id_nil,id_U11 => true
      | id_nil,id_isQid => true
      | id_nil,id_isNeList => true
      | id_nil,id_ok => true
      | id_nil,id_mark => true
      | id_nil,id_isPal => true
      | id_nil,id_U41 => true
      | id_nil,id_u => true
      | id_nil,id_U21 => true
      | id_nil,id_a => true
      | id_nil,id_U52 => true
      | id_nil,id___ => true
      | id_nil,id_U72 => true
      | id_nil,id_U31 => true
      | id_nil,id_o => true
      | id_nil,id_tt => true
      | id_nil,id_isNePal => true
      | id_nil,id_U51 => true
      | id_nil,id_top => true
      | id_nil,id_nil => true
      | id_nil,id_U81 => false
      | id_nil,id_U42 => false
      | id_nil,id_proper => false
      | id_nil,id_U22 => false
      | id_nil,id_e => false
      | id_nil,id_U61 => false
      | id_U81,id_active => true
      | id_U81,id_U71 => true
      | id_U81,id_isList => true
      | id_U81,id_i => true
      | id_U81,id_U11 => true
      | id_U81,id_isQid => true
      | id_U81,id_isNeList => true
      | id_U81,id_ok => true
      | id_U81,id_mark => true
      | id_U81,id_isPal => true
      | id_U81,id_U41 => true
      | id_U81,id_u => true
      | id_U81,id_U21 => true
      | id_U81,id_a => true
      | id_U81,id_U52 => true
      | id_U81,id___ => true
      | id_U81,id_U72 => true
      | id_U81,id_U31 => true
      | id_U81,id_o => true
      | id_U81,id_tt => true
      | id_U81,id_isNePal => true
      | id_U81,id_U51 => true
      | id_U81,id_top => true
      | id_U81,id_nil => true
      | id_U81,id_U81 => true
      | id_U81,id_U42 => false
      | id_U81,id_proper => false
      | id_U81,id_U22 => false
      | id_U81,id_e => false
      | id_U81,id_U61 => false
      | id_U42,id_active => true
      | id_U42,id_U71 => true
      | id_U42,id_isList => true
      | id_U42,id_i => true
      | id_U42,id_U11 => true
      | id_U42,id_isQid => true
      | id_U42,id_isNeList => true
      | id_U42,id_ok => true
      | id_U42,id_mark => true
      | id_U42,id_isPal => true
      | id_U42,id_U41 => true
      | id_U42,id_u => true
      | id_U42,id_U21 => true
      | id_U42,id_a => true
      | id_U42,id_U52 => true
      | id_U42,id___ => true
      | id_U42,id_U72 => true
      | id_U42,id_U31 => true
      | id_U42,id_o => true
      | id_U42,id_tt => true
      | id_U42,id_isNePal => true
      | id_U42,id_U51 => true
      | id_U42,id_top => true
      | id_U42,id_nil => true
      | id_U42,id_U81 => true
      | id_U42,id_U42 => true
      | id_U42,id_proper => false
      | id_U42,id_U22 => false
      | id_U42,id_e => false
      | id_U42,id_U61 => false
      | id_proper,id_active => true
      | id_proper,id_U71 => true
      | id_proper,id_isList => true
      | id_proper,id_i => true
      | id_proper,id_U11 => true
      | id_proper,id_isQid => true
      | id_proper,id_isNeList => true
      | id_proper,id_ok => true
      | id_proper,id_mark => true
      | id_proper,id_isPal => true
      | id_proper,id_U41 => true
      | id_proper,id_u => true
      | id_proper,id_U21 => true
      | id_proper,id_a => true
      | id_proper,id_U52 => true
      | id_proper,id___ => true
      | id_proper,id_U72 => true
      | id_proper,id_U31 => true
      | id_proper,id_o => true
      | id_proper,id_tt => true
      | id_proper,id_isNePal => true
      | id_proper,id_U51 => true
      | id_proper,id_top => true
      | id_proper,id_nil => true
      | id_proper,id_U81 => true
      | id_proper,id_U42 => true
      | id_proper,id_proper => true
      | id_proper,id_U22 => false
      | id_proper,id_e => false
      | id_proper,id_U61 => false
      | id_U22,id_active => true
      | id_U22,id_U71 => true
      | id_U22,id_isList => true
      | id_U22,id_i => true
      | id_U22,id_U11 => true
      | id_U22,id_isQid => true
      | id_U22,id_isNeList => true
      | id_U22,id_ok => true
      | id_U22,id_mark => true
      | id_U22,id_isPal => true
      | id_U22,id_U41 => true
      | id_U22,id_u => true
      | id_U22,id_U21 => true
      | id_U22,id_a => true
      | id_U22,id_U52 => true
      | id_U22,id___ => true
      | id_U22,id_U72 => true
      | id_U22,id_U31 => true
      | id_U22,id_o => true
      | id_U22,id_tt => true
      | id_U22,id_isNePal => true
      | id_U22,id_U51 => true
      | id_U22,id_top => true
      | id_U22,id_nil => true
      | id_U22,id_U81 => true
      | id_U22,id_U42 => true
      | id_U22,id_proper => true
      | id_U22,id_U22 => true
      | id_U22,id_e => false
      | id_U22,id_U61 => false
      | id_e,id_active => true
      | id_e,id_U71 => true
      | id_e,id_isList => true
      | id_e,id_i => true
      | id_e,id_U11 => true
      | id_e,id_isQid => true
      | id_e,id_isNeList => true
      | id_e,id_ok => true
      | id_e,id_mark => true
      | id_e,id_isPal => true
      | id_e,id_U41 => true
      | id_e,id_u => true
      | id_e,id_U21 => true
      | id_e,id_a => true
      | id_e,id_U52 => true
      | id_e,id___ => true
      | id_e,id_U72 => true
      | id_e,id_U31 => true
      | id_e,id_o => true
      | id_e,id_tt => true
      | id_e,id_isNePal => true
      | id_e,id_U51 => true
      | id_e,id_top => true
      | id_e,id_nil => true
      | id_e,id_U81 => true
      | id_e,id_U42 => true
      | id_e,id_proper => true
      | id_e,id_U22 => true
      | id_e,id_e => true
      | id_e,id_U61 => false
      | id_U61,id_active => true
      | id_U61,id_U71 => true
      | id_U61,id_isList => true
      | id_U61,id_i => true
      | id_U61,id_U11 => true
      | id_U61,id_isQid => true
      | id_U61,id_isNeList => true
      | id_U61,id_ok => true
      | id_U61,id_mark => true
      | id_U61,id_isPal => true
      | id_U61,id_U41 => true
      | id_U61,id_u => true
      | id_U61,id_U21 => true
      | id_U61,id_a => true
      | id_U61,id_U52 => true
      | id_U61,id___ => true
      | id_U61,id_U72 => true
      | id_U61,id_U31 => true
      | id_U61,id_o => true
      | id_U61,id_tt => true
      | id_U61,id_isNePal => true
      | id_U61,id_U51 => true
      | id_U61,id_top => true
      | id_U61,id_nil => true
      | id_U61,id_U81 => true
      | id_U61,id_U42 => true
      | id_U61,id_proper => true
      | id_U61,id_U22 => true
      | id_U61,id_e => true
      | id_U61,id_U61 => 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(__(__(X_,Y_),Z_)) -> mark(__(X_,__(Y_,Z_))) *)
   | R_xml_0_rule_0 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id___ ((algebra.Alg.Var 1)::(algebra.Alg.Term 
                   algebra.F.id___ ((algebra.Alg.Var 2)::
                   (algebra.Alg.Var 3)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id___ ((algebra.Alg.Term algebra.F.id___ 
      ((algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil))::
      (algebra.Alg.Var 3)::nil))::nil))
    (* active(__(X_,nil)) -> mark(X_) *)
   | R_xml_0_rule_1 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark 
                   ((algebra.Alg.Var 1)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id___ ((algebra.Alg.Var 1)::(algebra.Alg.Term 
      algebra.F.id_nil nil)::nil))::nil))
    (* active(__(nil,X_)) -> mark(X_) *)
   | R_xml_0_rule_2 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark 
                   ((algebra.Alg.Var 1)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id___ ((algebra.Alg.Term algebra.F.id_nil nil)::
      (algebra.Alg.Var 1)::nil))::nil))
    (* active(U11(tt)) -> mark(tt) *)
   | R_xml_0_rule_3 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_tt nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U11 ((algebra.Alg.Term algebra.F.id_tt nil)::nil))::nil))
    (* active(U21(tt,V2_)) -> mark(U22(isList(V2_))) *)
   | R_xml_0_rule_4 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U22 ((algebra.Alg.Term algebra.F.id_isList 
                   ((algebra.Alg.Var 4)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U21 ((algebra.Alg.Term algebra.F.id_tt nil)::
      (algebra.Alg.Var 4)::nil))::nil))
    (* active(U22(tt)) -> mark(tt) *)
   | R_xml_0_rule_5 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_tt nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U22 ((algebra.Alg.Term algebra.F.id_tt nil)::nil))::nil))
    (* active(U31(tt)) -> mark(tt) *)
   | R_xml_0_rule_6 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_tt nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U31 ((algebra.Alg.Term algebra.F.id_tt nil)::nil))::nil))
    (* active(U41(tt,V2_)) -> mark(U42(isNeList(V2_))) *)
   | R_xml_0_rule_7 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U42 ((algebra.Alg.Term algebra.F.id_isNeList 
                   ((algebra.Alg.Var 4)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U41 ((algebra.Alg.Term algebra.F.id_tt nil)::
      (algebra.Alg.Var 4)::nil))::nil))
    (* active(U42(tt)) -> mark(tt) *)
   | R_xml_0_rule_8 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_tt nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U42 ((algebra.Alg.Term algebra.F.id_tt nil)::nil))::nil))
    (* active(U51(tt,V2_)) -> mark(U52(isList(V2_))) *)
   | R_xml_0_rule_9 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U52 ((algebra.Alg.Term algebra.F.id_isList 
                   ((algebra.Alg.Var 4)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U51 ((algebra.Alg.Term algebra.F.id_tt nil)::
      (algebra.Alg.Var 4)::nil))::nil))
    (* active(U52(tt)) -> mark(tt) *)
   | R_xml_0_rule_10 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_tt nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U52 ((algebra.Alg.Term algebra.F.id_tt nil)::nil))::nil))
    (* active(U61(tt)) -> mark(tt) *)
   | R_xml_0_rule_11 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_tt nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U61 ((algebra.Alg.Term algebra.F.id_tt nil)::nil))::nil))
    (* active(U71(tt,P_)) -> mark(U72(isPal(P_))) *)
   | R_xml_0_rule_12 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U72 ((algebra.Alg.Term algebra.F.id_isPal 
                   ((algebra.Alg.Var 5)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U71 ((algebra.Alg.Term algebra.F.id_tt nil)::
      (algebra.Alg.Var 5)::nil))::nil))
    (* active(U72(tt)) -> mark(tt) *)
   | R_xml_0_rule_13 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_tt nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U72 ((algebra.Alg.Term algebra.F.id_tt nil)::nil))::nil))
    (* active(U81(tt)) -> mark(tt) *)
   | R_xml_0_rule_14 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_tt nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U81 ((algebra.Alg.Term algebra.F.id_tt nil)::nil))::nil))
    (* active(isList(V_)) -> mark(U11(isNeList(V_))) *)
   | R_xml_0_rule_15 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U11 ((algebra.Alg.Term algebra.F.id_isNeList 
                   ((algebra.Alg.Var 6)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_isList ((algebra.Alg.Var 6)::nil))::nil))
    (* active(isList(nil)) -> mark(tt) *)
   | R_xml_0_rule_16 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_tt nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_isList ((algebra.Alg.Term algebra.F.id_nil 
      nil)::nil))::nil))
   
    (* active(isList(__(V1_,V2_))) -> mark(U21(isList(V1_),V2_)) *)
   | R_xml_0_rule_17 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U21 ((algebra.Alg.Term algebra.F.id_isList 
                   ((algebra.Alg.Var 7)::nil))::
                   (algebra.Alg.Var 4)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_isList ((algebra.Alg.Term algebra.F.id___ 
      ((algebra.Alg.Var 7)::(algebra.Alg.Var 4)::nil))::nil))::nil))
    (* active(isNeList(V_)) -> mark(U31(isQid(V_))) *)
   | R_xml_0_rule_18 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U31 ((algebra.Alg.Term algebra.F.id_isQid 
                   ((algebra.Alg.Var 6)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_isNeList ((algebra.Alg.Var 6)::nil))::nil))
   
    (* active(isNeList(__(V1_,V2_))) -> mark(U41(isList(V1_),V2_)) *)
   | R_xml_0_rule_19 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U41 ((algebra.Alg.Term algebra.F.id_isList 
                   ((algebra.Alg.Var 7)::nil))::
                   (algebra.Alg.Var 4)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_isNeList ((algebra.Alg.Term algebra.F.id___ 
      ((algebra.Alg.Var 7)::(algebra.Alg.Var 4)::nil))::nil))::nil))
   
    (* active(isNeList(__(V1_,V2_))) -> mark(U51(isNeList(V1_),V2_)) *)
   | R_xml_0_rule_20 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U51 ((algebra.Alg.Term algebra.F.id_isNeList 
                   ((algebra.Alg.Var 7)::nil))::
                   (algebra.Alg.Var 4)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_isNeList ((algebra.Alg.Term algebra.F.id___ 
      ((algebra.Alg.Var 7)::(algebra.Alg.Var 4)::nil))::nil))::nil))
    (* active(isNePal(V_)) -> mark(U61(isQid(V_))) *)
   | R_xml_0_rule_21 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U61 ((algebra.Alg.Term algebra.F.id_isQid 
                   ((algebra.Alg.Var 6)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_isNePal ((algebra.Alg.Var 6)::nil))::nil))
   
    (* active(isNePal(__(I_,__(P_,I_)))) -> mark(U71(isQid(I_),P_)) *)
   | R_xml_0_rule_22 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U71 ((algebra.Alg.Term algebra.F.id_isQid 
                   ((algebra.Alg.Var 8)::nil))::
                   (algebra.Alg.Var 5)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_isNePal ((algebra.Alg.Term algebra.F.id___ 
      ((algebra.Alg.Var 8)::(algebra.Alg.Term algebra.F.id___ 
      ((algebra.Alg.Var 5)::(algebra.Alg.Var 8)::nil))::nil))::nil))::nil))
    (* active(isPal(V_)) -> mark(U81(isNePal(V_))) *)
   | R_xml_0_rule_23 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U81 ((algebra.Alg.Term algebra.F.id_isNePal 
                   ((algebra.Alg.Var 6)::nil))::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_isPal ((algebra.Alg.Var 6)::nil))::nil))
    (* active(isPal(nil)) -> mark(tt) *)
   | R_xml_0_rule_24 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_tt nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_isPal ((algebra.Alg.Term algebra.F.id_nil 
      nil)::nil))::nil))
    (* active(isQid(a)) -> mark(tt) *)
   | R_xml_0_rule_25 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_tt nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_isQid ((algebra.Alg.Term algebra.F.id_a nil)::nil))::nil))
    (* active(isQid(e)) -> mark(tt) *)
   | R_xml_0_rule_26 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_tt nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_isQid ((algebra.Alg.Term algebra.F.id_e nil)::nil))::nil))
    (* active(isQid(i)) -> mark(tt) *)
   | R_xml_0_rule_27 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_tt nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_isQid ((algebra.Alg.Term algebra.F.id_i nil)::nil))::nil))
    (* active(isQid(o)) -> mark(tt) *)
   | R_xml_0_rule_28 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_tt nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_isQid ((algebra.Alg.Term algebra.F.id_o nil)::nil))::nil))
    (* active(isQid(u)) -> mark(tt) *)
   | R_xml_0_rule_29 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_tt nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_isQid ((algebra.Alg.Term algebra.F.id_u nil)::nil))::nil))
    (* active(__(X1_,X2_)) -> __(active(X1_),X2_) *)
   | R_xml_0_rule_30 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Term 
                   algebra.F.id_active ((algebra.Alg.Var 9)::nil))::
                   (algebra.Alg.Var 10)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id___ ((algebra.Alg.Var 9)::
      (algebra.Alg.Var 10)::nil))::nil))
    (* active(__(X1_,X2_)) -> __(X1_,active(X2_)) *)
   | R_xml_0_rule_31 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 9)::
                   (algebra.Alg.Term algebra.F.id_active 
                   ((algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id___ ((algebra.Alg.Var 9)::
      (algebra.Alg.Var 10)::nil))::nil))
    (* active(U11(X_)) -> U11(active(X_)) *)
   | R_xml_0_rule_32 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Term 
                   algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U11 ((algebra.Alg.Var 1)::nil))::nil))
    (* active(U21(X1_,X2_)) -> U21(active(X1_),X2_) *)
   | R_xml_0_rule_33 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Term 
                   algebra.F.id_active ((algebra.Alg.Var 9)::nil))::
                   (algebra.Alg.Var 10)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U21 ((algebra.Alg.Var 9)::
      (algebra.Alg.Var 10)::nil))::nil))
    (* active(U22(X_)) -> U22(active(X_)) *)
   | R_xml_0_rule_34 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Term 
                   algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U22 ((algebra.Alg.Var 1)::nil))::nil))
    (* active(U31(X_)) -> U31(active(X_)) *)
   | R_xml_0_rule_35 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Term 
                   algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U31 ((algebra.Alg.Var 1)::nil))::nil))
    (* active(U41(X1_,X2_)) -> U41(active(X1_),X2_) *)
   | R_xml_0_rule_36 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Term 
                   algebra.F.id_active ((algebra.Alg.Var 9)::nil))::
                   (algebra.Alg.Var 10)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U41 ((algebra.Alg.Var 9)::
      (algebra.Alg.Var 10)::nil))::nil))
    (* active(U42(X_)) -> U42(active(X_)) *)
   | R_xml_0_rule_37 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Term 
                   algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U42 ((algebra.Alg.Var 1)::nil))::nil))
    (* active(U51(X1_,X2_)) -> U51(active(X1_),X2_) *)
   | R_xml_0_rule_38 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Term 
                   algebra.F.id_active ((algebra.Alg.Var 9)::nil))::
                   (algebra.Alg.Var 10)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U51 ((algebra.Alg.Var 9)::
      (algebra.Alg.Var 10)::nil))::nil))
    (* active(U52(X_)) -> U52(active(X_)) *)
   | R_xml_0_rule_39 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Term 
                   algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U52 ((algebra.Alg.Var 1)::nil))::nil))
    (* active(U61(X_)) -> U61(active(X_)) *)
   | R_xml_0_rule_40 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Term 
                   algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U61 ((algebra.Alg.Var 1)::nil))::nil))
    (* active(U71(X1_,X2_)) -> U71(active(X1_),X2_) *)
   | R_xml_0_rule_41 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Term 
                   algebra.F.id_active ((algebra.Alg.Var 9)::nil))::
                   (algebra.Alg.Var 10)::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U71 ((algebra.Alg.Var 9)::
      (algebra.Alg.Var 10)::nil))::nil))
    (* active(U72(X_)) -> U72(active(X_)) *)
   | R_xml_0_rule_42 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U72 ((algebra.Alg.Term 
                   algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U72 ((algebra.Alg.Var 1)::nil))::nil))
    (* active(U81(X_)) -> U81(active(X_)) *)
   | R_xml_0_rule_43 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U81 ((algebra.Alg.Term 
                   algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
      algebra.F.id_U81 ((algebra.Alg.Var 1)::nil))::nil))
    (* __(mark(X1_),X2_) -> mark(__(X1_,X2_)) *)
   | R_xml_0_rule_44 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id___ ((algebra.Alg.Var 9)::
                   (algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Var 10)::nil))
    (* __(X1_,mark(X2_)) -> mark(__(X1_,X2_)) *)
   | R_xml_0_rule_45 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id___ ((algebra.Alg.Var 9)::
                   (algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 9)::
      (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 10)::nil))::nil))
    (* U11(mark(X_)) -> mark(U11(X_)) *)
   | R_xml_0_rule_46 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U11 ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* U21(mark(X1_),X2_) -> mark(U21(X1_,X2_)) *)
   | R_xml_0_rule_47 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U21 ((algebra.Alg.Var 9)::
                   (algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Var 10)::nil))
    (* U22(mark(X_)) -> mark(U22(X_)) *)
   | R_xml_0_rule_48 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U22 ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* U31(mark(X_)) -> mark(U31(X_)) *)
   | R_xml_0_rule_49 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U31 ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* U41(mark(X1_),X2_) -> mark(U41(X1_,X2_)) *)
   | R_xml_0_rule_50 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U41 ((algebra.Alg.Var 9)::
                   (algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Var 10)::nil))
    (* U42(mark(X_)) -> mark(U42(X_)) *)
   | R_xml_0_rule_51 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U42 ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* U51(mark(X1_),X2_) -> mark(U51(X1_,X2_)) *)
   | R_xml_0_rule_52 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U51 ((algebra.Alg.Var 9)::
                   (algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Var 10)::nil))
    (* U52(mark(X_)) -> mark(U52(X_)) *)
   | R_xml_0_rule_53 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U52 ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* U61(mark(X_)) -> mark(U61(X_)) *)
   | R_xml_0_rule_54 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U61 ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* U71(mark(X1_),X2_) -> mark(U71(X1_,X2_)) *)
   | R_xml_0_rule_55 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U71 ((algebra.Alg.Var 9)::
                   (algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Var 10)::nil))
    (* U72(mark(X_)) -> mark(U72(X_)) *)
   | R_xml_0_rule_56 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U72 ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U72 ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* U81(mark(X_)) -> mark(U81(X_)) *)
   | R_xml_0_rule_57 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term 
                   algebra.F.id_U81 ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U81 ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 1)::nil))::nil))
   
    (* proper(__(X1_,X2_)) -> __(proper(X1_),proper(X2_)) *)
   | R_xml_0_rule_58 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 9)::nil))::
                   (algebra.Alg.Term algebra.F.id_proper 
                   ((algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id___ ((algebra.Alg.Var 9)::
      (algebra.Alg.Var 10)::nil))::nil))
    (* proper(nil) -> ok(nil) *)
   | R_xml_0_rule_59 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_nil nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_nil nil)::nil))
    (* proper(U11(X_)) -> U11(proper(X_)) *)
   | R_xml_0_rule_60 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_U11 ((algebra.Alg.Var 1)::nil))::nil))
    (* proper(tt) -> ok(tt) *)
   | R_xml_0_rule_61 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_tt nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_tt nil)::nil))
   
    (* proper(U21(X1_,X2_)) -> U21(proper(X1_),proper(X2_)) *)
   | R_xml_0_rule_62 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 9)::nil))::
                   (algebra.Alg.Term algebra.F.id_proper 
                   ((algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_U21 ((algebra.Alg.Var 9)::
      (algebra.Alg.Var 10)::nil))::nil))
    (* proper(U22(X_)) -> U22(proper(X_)) *)
   | R_xml_0_rule_63 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_U22 ((algebra.Alg.Var 1)::nil))::nil))
    (* proper(isList(X_)) -> isList(proper(X_)) *)
   | R_xml_0_rule_64 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_isList ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_isList ((algebra.Alg.Var 1)::nil))::nil))
    (* proper(U31(X_)) -> U31(proper(X_)) *)
   | R_xml_0_rule_65 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_U31 ((algebra.Alg.Var 1)::nil))::nil))
   
    (* proper(U41(X1_,X2_)) -> U41(proper(X1_),proper(X2_)) *)
   | R_xml_0_rule_66 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 9)::nil))::
                   (algebra.Alg.Term algebra.F.id_proper 
                   ((algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_U41 ((algebra.Alg.Var 9)::
      (algebra.Alg.Var 10)::nil))::nil))
    (* proper(U42(X_)) -> U42(proper(X_)) *)
   | R_xml_0_rule_67 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_U42 ((algebra.Alg.Var 1)::nil))::nil))
    (* proper(isNeList(X_)) -> isNeList(proper(X_)) *)
   | R_xml_0_rule_68 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_isNeList ((algebra.Alg.Var 1)::nil))::nil))
   
    (* proper(U51(X1_,X2_)) -> U51(proper(X1_),proper(X2_)) *)
   | R_xml_0_rule_69 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 9)::nil))::
                   (algebra.Alg.Term algebra.F.id_proper 
                   ((algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_U51 ((algebra.Alg.Var 9)::
      (algebra.Alg.Var 10)::nil))::nil))
    (* proper(U52(X_)) -> U52(proper(X_)) *)
   | R_xml_0_rule_70 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_U52 ((algebra.Alg.Var 1)::nil))::nil))
    (* proper(U61(X_)) -> U61(proper(X_)) *)
   | R_xml_0_rule_71 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_U61 ((algebra.Alg.Var 1)::nil))::nil))
   
    (* proper(U71(X1_,X2_)) -> U71(proper(X1_),proper(X2_)) *)
   | R_xml_0_rule_72 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 9)::nil))::
                   (algebra.Alg.Term algebra.F.id_proper 
                   ((algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_U71 ((algebra.Alg.Var 9)::
      (algebra.Alg.Var 10)::nil))::nil))
    (* proper(U72(X_)) -> U72(proper(X_)) *)
   | R_xml_0_rule_73 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U72 ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_U72 ((algebra.Alg.Var 1)::nil))::nil))
    (* proper(isPal(X_)) -> isPal(proper(X_)) *)
   | R_xml_0_rule_74 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_isPal ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_isPal ((algebra.Alg.Var 1)::nil))::nil))
    (* proper(U81(X_)) -> U81(proper(X_)) *)
   | R_xml_0_rule_75 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_U81 ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_U81 ((algebra.Alg.Var 1)::nil))::nil))
    (* proper(isQid(X_)) -> isQid(proper(X_)) *)
   | R_xml_0_rule_76 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_isQid ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_isQid ((algebra.Alg.Var 1)::nil))::nil))
    (* proper(isNePal(X_)) -> isNePal(proper(X_)) *)
   | R_xml_0_rule_77 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_isNePal ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
      algebra.F.id_isNePal ((algebra.Alg.Var 1)::nil))::nil))
    (* proper(a) -> ok(a) *)
   | R_xml_0_rule_78 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_a nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term algebra.F.id_a 
      nil)::nil))
    (* proper(e) -> ok(e) *)
   | R_xml_0_rule_79 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_e nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term algebra.F.id_e 
      nil)::nil))
    (* proper(i) -> ok(i) *)
   | R_xml_0_rule_80 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_i nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term algebra.F.id_i 
      nil)::nil))
    (* proper(o) -> ok(o) *)
   | R_xml_0_rule_81 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_o nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term algebra.F.id_o 
      nil)::nil))
    (* proper(u) -> ok(u) *)
   | R_xml_0_rule_82 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_u nil)::nil)) 
     (algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term algebra.F.id_u 
      nil)::nil))
    (* __(ok(X1_),ok(X2_)) -> ok(__(X1_,X2_)) *)
   | R_xml_0_rule_83 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id___ ((algebra.Alg.Var 9)::
                   (algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 10)::nil))::nil))
    (* U11(ok(X_)) -> ok(U11(X_)) *)
   | R_xml_0_rule_84 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_U11 ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* U21(ok(X1_),ok(X2_)) -> ok(U21(X1_,X2_)) *)
   | R_xml_0_rule_85 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_U21 ((algebra.Alg.Var 9)::
                   (algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 10)::nil))::nil))
    (* U22(ok(X_)) -> ok(U22(X_)) *)
   | R_xml_0_rule_86 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_U22 ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* isList(ok(X_)) -> ok(isList(X_)) *)
   | R_xml_0_rule_87 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_isList ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_isList ((algebra.Alg.Term 
      algebra.F.id_ok ((algebra.Alg.Var 1)::nil))::nil))
    (* U31(ok(X_)) -> ok(U31(X_)) *)
   | R_xml_0_rule_88 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_U31 ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* U41(ok(X1_),ok(X2_)) -> ok(U41(X1_,X2_)) *)
   | R_xml_0_rule_89 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_U41 ((algebra.Alg.Var 9)::
                   (algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 10)::nil))::nil))
    (* U42(ok(X_)) -> ok(U42(X_)) *)
   | R_xml_0_rule_90 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_U42 ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* isNeList(ok(X_)) -> ok(isNeList(X_)) *)
   | R_xml_0_rule_91 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_isNeList ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Term 
      algebra.F.id_ok ((algebra.Alg.Var 1)::nil))::nil))
    (* U51(ok(X1_),ok(X2_)) -> ok(U51(X1_,X2_)) *)
   | R_xml_0_rule_92 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_U51 ((algebra.Alg.Var 9)::
                   (algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 10)::nil))::nil))
    (* U52(ok(X_)) -> ok(U52(X_)) *)
   | R_xml_0_rule_93 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_U52 ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* U61(ok(X_)) -> ok(U61(X_)) *)
   | R_xml_0_rule_94 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_U61 ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* U71(ok(X1_),ok(X2_)) -> ok(U71(X1_,X2_)) *)
   | R_xml_0_rule_95 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_U71 ((algebra.Alg.Var 9)::
                   (algebra.Alg.Var 10)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 10)::nil))::nil))
    (* U72(ok(X_)) -> ok(U72(X_)) *)
   | R_xml_0_rule_96 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_U72 ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U72 ((algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* isPal(ok(X_)) -> ok(isPal(X_)) *)
   | R_xml_0_rule_97 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_isPal ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_isPal ((algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* U81(ok(X_)) -> ok(U81(X_)) *)
   | R_xml_0_rule_98 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_U81 ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_U81 ((algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* isQid(ok(X_)) -> ok(isQid(X_)) *)
   | R_xml_0_rule_99 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_isQid ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_isQid ((algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* isNePal(ok(X_)) -> ok(isNePal(X_)) *)
   | R_xml_0_rule_100 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
                   algebra.F.id_isNePal ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_isNePal ((algebra.Alg.Term 
      algebra.F.id_ok ((algebra.Alg.Var 1)::nil))::nil))
    (* top(mark(X_)) -> top(proper(X_)) *)
   | R_xml_0_rule_101 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_top ((algebra.Alg.Term 
                   algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_top ((algebra.Alg.Term algebra.F.id_mark 
      ((algebra.Alg.Var 1)::nil))::nil))
    (* top(ok(X_)) -> top(active(X_)) *)
   | R_xml_0_rule_102 :
    R_xml_0_rules (algebra.Alg.Term algebra.F.id_top ((algebra.Alg.Term 
                   algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)) 
     (algebra.Alg.Term algebra.F.id_top ((algebra.Alg.Term algebra.F.id_ok 
      ((algebra.Alg.Var 1)::nil))::nil))
 .
 
 
 Definition R_xml_0_rule_as_list_0  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term algebra.F.id___ 
     ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 1)::
     (algebra.Alg.Var 2)::nil))::(algebra.Alg.Var 3)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id___ 
     ((algebra.Alg.Var 1)::(algebra.Alg.Term algebra.F.id___ 
     ((algebra.Alg.Var 2)::(algebra.Alg.Var 3)::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___ 
     ((algebra.Alg.Var 1)::(algebra.Alg.Term algebra.F.id_nil 
     nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::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___ 
     ((algebra.Alg.Term algebra.F.id_nil nil)::
     (algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::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_U11 ((algebra.Alg.Term algebra.F.id_tt nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt 
     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_U21 ((algebra.Alg.Term algebra.F.id_tt nil)::
     (algebra.Alg.Var 4)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U22 
     ((algebra.Alg.Term algebra.F.id_isList 
     ((algebra.Alg.Var 4)::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_U22 ((algebra.Alg.Term algebra.F.id_tt nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt 
     nil)::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_U31 ((algebra.Alg.Term algebra.F.id_tt nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt 
     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_U41 ((algebra.Alg.Term algebra.F.id_tt nil)::
     (algebra.Alg.Var 4)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U42 
     ((algebra.Alg.Term algebra.F.id_isNeList 
     ((algebra.Alg.Var 4)::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_U42 ((algebra.Alg.Term algebra.F.id_tt nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt 
     nil)::nil)))::R_xml_0_rule_as_list_7.
 
 
 Definition R_xml_0_rule_as_list_9  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U51 ((algebra.Alg.Term algebra.F.id_tt nil)::
     (algebra.Alg.Var 4)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U52 
     ((algebra.Alg.Term algebra.F.id_isList 
     ((algebra.Alg.Var 4)::nil))::nil))::nil)))::R_xml_0_rule_as_list_8.
 
 
 Definition R_xml_0_rule_as_list_10  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U52 ((algebra.Alg.Term algebra.F.id_tt nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt 
     nil)::nil)))::R_xml_0_rule_as_list_9.
 
 
 Definition R_xml_0_rule_as_list_11  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U61 ((algebra.Alg.Term algebra.F.id_tt nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt 
     nil)::nil)))::R_xml_0_rule_as_list_10.
 
 
 Definition R_xml_0_rule_as_list_12  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U71 ((algebra.Alg.Term algebra.F.id_tt nil)::
     (algebra.Alg.Var 5)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U72 
     ((algebra.Alg.Term algebra.F.id_isPal 
     ((algebra.Alg.Var 5)::nil))::nil))::nil)))::R_xml_0_rule_as_list_11.
 
 
 Definition R_xml_0_rule_as_list_13  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U72 ((algebra.Alg.Term algebra.F.id_tt nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt 
     nil)::nil)))::R_xml_0_rule_as_list_12.
 
 
 Definition R_xml_0_rule_as_list_14  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U81 ((algebra.Alg.Term algebra.F.id_tt nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt 
     nil)::nil)))::R_xml_0_rule_as_list_13.
 
 
 Definition R_xml_0_rule_as_list_15  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_isList ((algebra.Alg.Var 6)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U11 
     ((algebra.Alg.Term algebra.F.id_isNeList 
     ((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_active ((algebra.Alg.Term 
     algebra.F.id_isList ((algebra.Alg.Term algebra.F.id_nil 
     nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt 
     nil)::nil)))::R_xml_0_rule_as_list_15.
 
 
 Definition R_xml_0_rule_as_list_17  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_isList ((algebra.Alg.Term algebra.F.id___ 
     ((algebra.Alg.Var 7)::(algebra.Alg.Var 4)::nil))::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U21 
     ((algebra.Alg.Term algebra.F.id_isList ((algebra.Alg.Var 7)::nil))::
     (algebra.Alg.Var 4)::nil))::nil)))::R_xml_0_rule_as_list_16.
 
 
 Definition R_xml_0_rule_as_list_18  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_isNeList ((algebra.Alg.Var 6)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U31 
     ((algebra.Alg.Term algebra.F.id_isQid 
     ((algebra.Alg.Var 6)::nil))::nil))::nil)))::R_xml_0_rule_as_list_17.
 
 
 Definition R_xml_0_rule_as_list_19  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_isNeList ((algebra.Alg.Term algebra.F.id___ 
     ((algebra.Alg.Var 7)::(algebra.Alg.Var 4)::nil))::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U41 
     ((algebra.Alg.Term algebra.F.id_isList ((algebra.Alg.Var 7)::nil))::
     (algebra.Alg.Var 4)::nil))::nil)))::R_xml_0_rule_as_list_18.
 
 
 Definition R_xml_0_rule_as_list_20  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_isNeList ((algebra.Alg.Term algebra.F.id___ 
     ((algebra.Alg.Var 7)::(algebra.Alg.Var 4)::nil))::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U51 
     ((algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Var 7)::nil))::
     (algebra.Alg.Var 4)::nil))::nil)))::R_xml_0_rule_as_list_19.
 
 
 Definition R_xml_0_rule_as_list_21  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_isNePal ((algebra.Alg.Var 6)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U61 
     ((algebra.Alg.Term algebra.F.id_isQid 
     ((algebra.Alg.Var 6)::nil))::nil))::nil)))::R_xml_0_rule_as_list_20.
 
 
 Definition R_xml_0_rule_as_list_22  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_isNePal ((algebra.Alg.Term algebra.F.id___ 
     ((algebra.Alg.Var 8)::(algebra.Alg.Term algebra.F.id___ 
     ((algebra.Alg.Var 5)::(algebra.Alg.Var 8)::nil))::nil))::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U71 
     ((algebra.Alg.Term algebra.F.id_isQid ((algebra.Alg.Var 8)::nil))::
     (algebra.Alg.Var 5)::nil))::nil)))::R_xml_0_rule_as_list_21.
 
 
 Definition R_xml_0_rule_as_list_23  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_isPal ((algebra.Alg.Var 6)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U81 
     ((algebra.Alg.Term algebra.F.id_isNePal 
     ((algebra.Alg.Var 6)::nil))::nil))::nil)))::R_xml_0_rule_as_list_22.
 
 
 Definition R_xml_0_rule_as_list_24  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_isPal ((algebra.Alg.Term algebra.F.id_nil 
     nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt 
     nil)::nil)))::R_xml_0_rule_as_list_23.
 
 
 Definition R_xml_0_rule_as_list_25  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_isQid ((algebra.Alg.Term algebra.F.id_a nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt 
     nil)::nil)))::R_xml_0_rule_as_list_24.
 
 
 Definition R_xml_0_rule_as_list_26  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_isQid ((algebra.Alg.Term algebra.F.id_e nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt 
     nil)::nil)))::R_xml_0_rule_as_list_25.
 
 
 Definition R_xml_0_rule_as_list_27  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_isQid ((algebra.Alg.Term algebra.F.id_i nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt 
     nil)::nil)))::R_xml_0_rule_as_list_26.
 
 
 Definition R_xml_0_rule_as_list_28  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_isQid ((algebra.Alg.Term algebra.F.id_o nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt 
     nil)::nil)))::R_xml_0_rule_as_list_27.
 
 
 Definition R_xml_0_rule_as_list_29  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_isQid ((algebra.Alg.Term algebra.F.id_u nil)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt 
     nil)::nil)))::R_xml_0_rule_as_list_28.
 
 
 Definition R_xml_0_rule_as_list_30  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term algebra.F.id___ 
     ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Term algebra.F.id_active 
     ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Var 10)::nil)))::
    R_xml_0_rule_as_list_29.
 
 
 Definition R_xml_0_rule_as_list_31  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term algebra.F.id___ 
     ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 9)::
     (algebra.Alg.Term algebra.F.id_active 
     ((algebra.Alg.Var 10)::nil))::nil)))::R_xml_0_rule_as_list_30.
 
 
 Definition R_xml_0_rule_as_list_32  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U11 ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_31.
 
 
 Definition R_xml_0_rule_as_list_33  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U21 ((algebra.Alg.Var 9)::
     (algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 9)::nil))::
     (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_32.
 
 
 Definition R_xml_0_rule_as_list_34  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U22 ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_33.
 
 
 Definition R_xml_0_rule_as_list_35  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U31 ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_34.
 
 
 Definition R_xml_0_rule_as_list_36  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U41 ((algebra.Alg.Var 9)::
     (algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 9)::nil))::
     (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_35.
 
 
 Definition R_xml_0_rule_as_list_37  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U42 ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_36.
 
 
 Definition R_xml_0_rule_as_list_38  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U51 ((algebra.Alg.Var 9)::
     (algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 9)::nil))::
     (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_37.
 
 
 Definition R_xml_0_rule_as_list_39  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U52 ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_38.
 
 
 Definition R_xml_0_rule_as_list_40  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U61 ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_39.
 
 
 Definition R_xml_0_rule_as_list_41  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U71 ((algebra.Alg.Var 9)::
     (algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 9)::nil))::
     (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_40.
 
 
 Definition R_xml_0_rule_as_list_42  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U72 ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U72 ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_41.
 
 
 Definition R_xml_0_rule_as_list_43  := 
   ((algebra.Alg.Term algebra.F.id_active ((algebra.Alg.Term 
     algebra.F.id_U81 ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U81 ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_42.
 
 
 Definition R_xml_0_rule_as_list_44  := 
   ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Var 10)::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id___ 
     ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)))::
    R_xml_0_rule_as_list_43.
 
 
 Definition R_xml_0_rule_as_list_45  := 
   ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 9)::
     (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id___ 
     ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)))::
    R_xml_0_rule_as_list_44.
 
 
 Definition R_xml_0_rule_as_list_46  := 
   ((algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U11 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_45.
 
 
 Definition R_xml_0_rule_as_list_47  := 
   ((algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Var 10)::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U21 
     ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)))::
    R_xml_0_rule_as_list_46.
 
 
 Definition R_xml_0_rule_as_list_48  := 
   ((algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U22 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_47.
 
 
 Definition R_xml_0_rule_as_list_49  := 
   ((algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U31 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_48.
 
 
 Definition R_xml_0_rule_as_list_50  := 
   ((algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Var 10)::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U41 
     ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)))::
    R_xml_0_rule_as_list_49.
 
 
 Definition R_xml_0_rule_as_list_51  := 
   ((algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U42 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_50.
 
 
 Definition R_xml_0_rule_as_list_52  := 
   ((algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Var 10)::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U51 
     ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)))::
    R_xml_0_rule_as_list_51.
 
 
 Definition R_xml_0_rule_as_list_53  := 
   ((algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U52 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_52.
 
 
 Definition R_xml_0_rule_as_list_54  := 
   ((algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U61 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_53.
 
 
 Definition R_xml_0_rule_as_list_55  := 
   ((algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Var 10)::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U71 
     ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)))::
    R_xml_0_rule_as_list_54.
 
 
 Definition R_xml_0_rule_as_list_56  := 
   ((algebra.Alg.Term algebra.F.id_U72 ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U72 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_55.
 
 
 Definition R_xml_0_rule_as_list_57  := 
   ((algebra.Alg.Term algebra.F.id_U81 ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U81 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_56.
 
 
 Definition R_xml_0_rule_as_list_58  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term algebra.F.id___ 
     ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Term algebra.F.id_proper 
     ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_proper 
     ((algebra.Alg.Var 10)::nil))::nil)))::R_xml_0_rule_as_list_57.
 
 
 Definition R_xml_0_rule_as_list_59  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_nil nil)::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_nil 
     nil)::nil)))::R_xml_0_rule_as_list_58.
 
 
 Definition R_xml_0_rule_as_list_60  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_U11 ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_59.
 
 
 Definition R_xml_0_rule_as_list_61  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term algebra.F.id_tt 
     nil)::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_tt 
     nil)::nil)))::R_xml_0_rule_as_list_60.
 
 
 Definition R_xml_0_rule_as_list_62  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_U21 ((algebra.Alg.Var 9)::
     (algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 10)::nil))::nil)))::
    R_xml_0_rule_as_list_61.
 
 
 Definition R_xml_0_rule_as_list_63  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_U22 ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_62.
 
 
 Definition R_xml_0_rule_as_list_64  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_isList ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_isList ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_63.
 
 
 Definition R_xml_0_rule_as_list_65  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_U31 ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_64.
 
 
 Definition R_xml_0_rule_as_list_66  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_U41 ((algebra.Alg.Var 9)::
     (algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 10)::nil))::nil)))::
    R_xml_0_rule_as_list_65.
 
 
 Definition R_xml_0_rule_as_list_67  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_U42 ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_66.
 
 
 Definition R_xml_0_rule_as_list_68  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_isNeList ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_67.
 
 
 Definition R_xml_0_rule_as_list_69  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_U51 ((algebra.Alg.Var 9)::
     (algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 10)::nil))::nil)))::
    R_xml_0_rule_as_list_68.
 
 
 Definition R_xml_0_rule_as_list_70  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_U52 ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_69.
 
 
 Definition R_xml_0_rule_as_list_71  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_U61 ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_70.
 
 
 Definition R_xml_0_rule_as_list_72  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_U71 ((algebra.Alg.Var 9)::
     (algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 10)::nil))::nil)))::
    R_xml_0_rule_as_list_71.
 
 
 Definition R_xml_0_rule_as_list_73  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_U72 ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U72 ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_72.
 
 
 Definition R_xml_0_rule_as_list_74  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_isPal ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_isPal ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_73.
 
 
 Definition R_xml_0_rule_as_list_75  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_U81 ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_U81 ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_74.
 
 
 Definition R_xml_0_rule_as_list_76  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_isQid ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_isQid ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_75.
 
 
 Definition R_xml_0_rule_as_list_77  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term 
     algebra.F.id_isNePal ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_isNePal ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_76.
 
 
 Definition R_xml_0_rule_as_list_78  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term algebra.F.id_a 
     nil)::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_a 
     nil)::nil)))::R_xml_0_rule_as_list_77.
 
 
 Definition R_xml_0_rule_as_list_79  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term algebra.F.id_e 
     nil)::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_e 
     nil)::nil)))::R_xml_0_rule_as_list_78.
 
 
 Definition R_xml_0_rule_as_list_80  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term algebra.F.id_i 
     nil)::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_i 
     nil)::nil)))::R_xml_0_rule_as_list_79.
 
 
 Definition R_xml_0_rule_as_list_81  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term algebra.F.id_o 
     nil)::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_o 
     nil)::nil)))::R_xml_0_rule_as_list_80.
 
 
 Definition R_xml_0_rule_as_list_82  := 
   ((algebra.Alg.Term algebra.F.id_proper ((algebra.Alg.Term algebra.F.id_u 
     nil)::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_u 
     nil)::nil)))::R_xml_0_rule_as_list_81.
 
 
 Definition R_xml_0_rule_as_list_83  := 
   ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id___ 
     ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)))::
    R_xml_0_rule_as_list_82.
 
 
 Definition R_xml_0_rule_as_list_84  := 
   ((algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_U11 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_83.
 
 
 Definition R_xml_0_rule_as_list_85  := 
   ((algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_U21 
     ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)))::
    R_xml_0_rule_as_list_84.
 
 
 Definition R_xml_0_rule_as_list_86  := 
   ((algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_U22 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_85.
 
 
 Definition R_xml_0_rule_as_list_87  := 
   ((algebra.Alg.Term algebra.F.id_isList ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_isList 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_86.
 
 
 Definition R_xml_0_rule_as_list_88  := 
   ((algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_U31 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_87.
 
 
 Definition R_xml_0_rule_as_list_89  := 
   ((algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_U41 
     ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)))::
    R_xml_0_rule_as_list_88.
 
 
 Definition R_xml_0_rule_as_list_90  := 
   ((algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_U42 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_89.
 
 
 Definition R_xml_0_rule_as_list_91  := 
   ((algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Term 
     algebra.F.id_ok ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
     algebra.F.id_isNeList ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_90.
 
 
 Definition R_xml_0_rule_as_list_92  := 
   ((algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_U51 
     ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)))::
    R_xml_0_rule_as_list_91.
 
 
 Definition R_xml_0_rule_as_list_93  := 
   ((algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_U52 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_92.
 
 
 Definition R_xml_0_rule_as_list_94  := 
   ((algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_U61 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_93.
 
 
 Definition R_xml_0_rule_as_list_95  := 
   ((algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 10)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_U71 
     ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)))::
    R_xml_0_rule_as_list_94.
 
 
 Definition R_xml_0_rule_as_list_96  := 
   ((algebra.Alg.Term algebra.F.id_U72 ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_U72 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_95.
 
 
 Definition R_xml_0_rule_as_list_97  := 
   ((algebra.Alg.Term algebra.F.id_isPal ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_isPal 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_96.
 
 
 Definition R_xml_0_rule_as_list_98  := 
   ((algebra.Alg.Term algebra.F.id_U81 ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_U81 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_97.
 
 
 Definition R_xml_0_rule_as_list_99  := 
   ((algebra.Alg.Term algebra.F.id_isQid ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term algebra.F.id_isQid 
     ((algebra.Alg.Var 1)::nil))::nil)))::R_xml_0_rule_as_list_98.
 
 
 Definition R_xml_0_rule_as_list_100  := 
   ((algebra.Alg.Term algebra.F.id_isNePal ((algebra.Alg.Term 
     algebra.F.id_ok ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_ok ((algebra.Alg.Term 
     algebra.F.id_isNePal ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_99.
 
 
 Definition R_xml_0_rule_as_list_101  := 
   ((algebra.Alg.Term algebra.F.id_top ((algebra.Alg.Term algebra.F.id_mark 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_top ((algebra.Alg.Term 
     algebra.F.id_proper ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_100.
 
 
 Definition R_xml_0_rule_as_list_102  := 
   ((algebra.Alg.Term algebra.F.id_top ((algebra.Alg.Term algebra.F.id_ok 
     ((algebra.Alg.Var 1)::nil))::nil)),
    (algebra.Alg.Term algebra.F.id_top ((algebra.Alg.Term 
     algebra.F.id_active ((algebra.Alg.Var 1)::nil))::nil)))::
    R_xml_0_rule_as_list_101.
 
 Definition R_xml_0_rule_as_list  := R_xml_0_rule_as_list_102.
 
 
 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 103.
   injection H;intros ;subst;constructor 102.
   injection H;intros ;subst;constructor 101.
   injection H;intros ;subst;constructor 100.
   injection H;intros ;subst;constructor 99.
   injection H;intros ;subst;constructor 98.
   injection H;intros ;subst;constructor 97.
   injection H;intros ;subst;constructor 96.
   injection H;intros ;subst;constructor 95.
   injection H;intros ;subst;constructor 94.
   injection H;intros ;subst;constructor 93.
   injection H;intros ;subst;constructor 92.
   injection H;intros ;subst;constructor 91.
   injection H;intros ;subst;constructor 90.
   injection H;intros ;subst;constructor 89.
   injection H;intros ;subst;constructor 88.
   injection H;intros ;subst;constructor 87.
   injection H;intros ;subst;constructor 86.
   injection H;intros ;subst;constructor 85.
   injection H;intros ;subst;constructor 84.
   injection H;intros ;subst;constructor 83.
   injection H;intros ;subst;constructor 82.
   injection H;intros ;subst;constructor 81.
   injection H;intros ;subst;constructor 80.
   rewrite  <- or_ext_generated.or25_equiv in H|-.
   case H;clear H;intros H.
   injection H;intros ;subst;constructor 79.
   injection H;intros ;subst;constructor 78.
   injection H;intros ;subst;constructor 77.
   injection H;intros ;subst;constructor 76.
   injection H;intros ;subst;constructor 75.
   injection H;intros ;subst;constructor 74.
   injection H;intros ;subst;constructor 73.
   injection H;intros ;subst;constructor 72.
   injection H;intros ;subst;constructor 71.
   injection H;intros ;subst;constructor 70.
   injection H;intros ;subst;constructor 69.
   injection H;intros ;subst;constructor 68.
   injection H;intros ;subst;constructor 67.
   injection H;intros ;subst;constructor 66.
   injection H;intros ;subst;constructor 65.
   injection H;intros ;subst;constructor 64.
   injection H;intros ;subst;constructor 63.
   injection H;intros ;subst;constructor 62.
   injection H;intros ;subst;constructor 61.
   injection H;intros ;subst;constructor 60.
   injection H;intros ;subst;constructor 59.
   injection H;intros ;subst;constructor 58.
   injection H;intros ;subst;constructor 57.
   injection H;intros ;subst;constructor 56.
   rewrite  <- or_ext_generated.or25_equiv in H|-.
   case H;clear H;intros H.
   injection H;intros ;subst;constructor 55.
   injection H;intros ;subst;constructor 54.
   injection H;intros ;subst;constructor 53.
   injection H;intros ;subst;constructor 52.
   injection H;intros ;subst;constructor 51.
   injection H;intros ;subst;constructor 50.
   injection H;intros ;subst;constructor 49.
   injection H;intros ;subst;constructor 48.
   injection H;intros ;subst;constructor 47.
   injection H;intros ;subst;constructor 46.
   injection H;intros ;subst;constructor 45.
   injection H;intros ;subst;constructor 44.
   injection H;intros ;subst;constructor 43.
   injection H;intros ;subst;constructor 42.
   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.
   rewrite  <- or_ext_generated.or25_equiv in H|-.
   case H;clear H;intros H.
   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.
   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.
   rewrite  <- or_ext_generated.or8_equiv in H|-.
   case H;clear H;intros H.
   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_9 (x12 x13 x14 x15 x16 x17 x18 x19 x20:Prop) :
  Prop := 
   | conj_9 :
    x12->x13->x14->x15->x16->x17->x18->x19->x20->
     and_9 x12 x13 x14 x15 x16 x17 x18 x19 x20
 .
 
 
 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_9 (t = (algebra.Alg.Term algebra.F.id_i nil) ->
           t' = (algebra.Alg.Term algebra.F.id_i nil)) 
     (forall x13, 
      t = (algebra.Alg.Term algebra.F.id_ok (x13::nil)) ->
       exists x12,
         t' = (algebra.Alg.Term algebra.F.id_ok (x12::nil))/\ 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
           x12 x13)) 
     (forall x13, 
      t = (algebra.Alg.Term algebra.F.id_mark (x13::nil)) ->
       exists x12,
         t' = (algebra.Alg.Term algebra.F.id_mark (x12::nil))/\ 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
           x12 x13)) 
     (t = (algebra.Alg.Term algebra.F.id_u nil) ->
      t' = (algebra.Alg.Term algebra.F.id_u nil)) 
     (t = (algebra.Alg.Term algebra.F.id_a nil) ->
      t' = (algebra.Alg.Term algebra.F.id_a nil)) 
     (t = (algebra.Alg.Term algebra.F.id_o nil) ->
      t' = (algebra.Alg.Term algebra.F.id_o nil)) 
     (t = (algebra.Alg.Term algebra.F.id_tt nil) ->
      t' = (algebra.Alg.Term algebra.F.id_tt nil)) 
     (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_e nil) ->
      t' = (algebra.Alg.Term algebra.F.id_e 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 x13 H;exists x13;intuition;constructor 1.
   intros x13 H;exists x13;intuition;constructor 1.
   intros H;intuition;constructor 1.
   intros H;intuition;constructor 1.
   intros H;intuition;constructor 1.
   intros H;intuition;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_i H_id_ok H_id_mark H_id_u H_id_a H_id_o H_id_tt H_id_nil H_id_e].
   constructor.
   
   clear H_id_ok H_id_mark H_id_u H_id_a H_id_o H_id_tt H_id_nil H_id_e;
    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_i H_id_mark H_id_u H_id_a H_id_o H_id_tt H_id_nil H_id_e;
    intros x13 H;injection H;clear H;intros ;subst;
    repeat (
    match goal with
      | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ =>
       inversion H;clear H;subst
      end
    ).
   
   match goal with
     | H:algebra.EQT.one_step _ ?y x13 |- _ =>
      destruct (H_id_ok y (refl_equal _)) as [x12];intros ;intuition;
       exists x12;intuition;eapply closure_extension.refl_trans_clos_R;
       eassumption
     end
   .
   
   clear H_id_i H_id_ok H_id_u H_id_a H_id_o H_id_tt H_id_nil H_id_e;
    intros x13 H;injection H;clear H;intros ;subst;
    repeat (
    match goal with
      | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ =>
       inversion H;clear H;subst
      end
    ).
   
   match goal with
     | H:algebra.EQT.one_step _ ?y x13 |- _ =>
      destruct (H_id_mark y (refl_equal _)) as [x12];intros ;intuition;
       exists x12;intuition;eapply closure_extension.refl_trans_clos_R;
       eassumption
     end
   .
   
   clear H_id_i H_id_ok H_id_mark H_id_a H_id_o H_id_tt H_id_nil H_id_e;
    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_i H_id_ok H_id_mark H_id_u H_id_o H_id_tt H_id_nil H_id_e;
    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_i H_id_ok H_id_mark H_id_u H_id_a H_id_tt H_id_nil H_id_e;
    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_i H_id_ok H_id_mark H_id_u H_id_a H_id_o H_id_nil H_id_e;
    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_i H_id_ok H_id_mark H_id_u H_id_a H_id_o H_id_tt H_id_e;
    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_i H_id_ok H_id_mark H_id_u H_id_a H_id_o H_id_tt 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_i_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_i nil) ->
     t' = (algebra.Alg.Term algebra.F.id_i nil).
 Proof.
   intros t t' H.
   destruct (are_constuctors_of_R_xml_0 H).
   assumption.
 Qed.
 
 
 Lemma id_ok_is_R_xml_0_constructor :
  forall t t', 
   (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) ->
    forall x13, 
     t = (algebra.Alg.Term algebra.F.id_ok (x13::nil)) ->
      exists x12,
        t' = (algebra.Alg.Term algebra.F.id_ok (x12::nil))/\ 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x12 x13).
 Proof.
   intros t t' H.
   destruct (are_constuctors_of_R_xml_0 H).
   assumption.
 Qed.
 
 
 Lemma id_mark_is_R_xml_0_constructor :
  forall t t', 
   (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) ->
    forall x13, 
     t = (algebra.Alg.Term algebra.F.id_mark (x13::nil)) ->
      exists x12,
        t' = (algebra.Alg.Term algebra.F.id_mark (x12::nil))/\ 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x12 x13).
 Proof.
   intros t t' H.
   destruct (are_constuctors_of_R_xml_0 H).
   assumption.
 Qed.
 
 
 Lemma id_u_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_u nil) ->
     t' = (algebra.Alg.Term algebra.F.id_u nil).
 Proof.
   intros t t' H.
   destruct (are_constuctors_of_R_xml_0 H).
   assumption.
 Qed.
 
 
 Lemma id_a_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_a nil) ->
     t' = (algebra.Alg.Term algebra.F.id_a nil).
 Proof.
   intros t t' H.
   destruct (are_constuctors_of_R_xml_0 H).
   assumption.
 Qed.
 
 
 Lemma id_o_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_o nil) ->
     t' = (algebra.Alg.Term algebra.F.id_o nil).
 Proof.
   intros t t' H.
   destruct (are_constuctors_of_R_xml_0 H).
   assumption.
 Qed.
 
 
 Lemma id_tt_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_tt nil) ->
     t' = (algebra.Alg.Term algebra.F.id_tt nil).
 Proof.
   intros t t' H.
   destruct (are_constuctors_of_R_xml_0 H).
   assumption.
 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_e_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_e nil) ->
     t' = (algebra.Alg.Term algebra.F.id_e 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_i nil) |- _ =>
     let Heq := fresh "Heq" in 
      (set (Heq:=id_i_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_ok (?x12::nil)) |- _ =>
     let x12 := fresh "x" in 
      (let Heq := fresh "Heq" in 
        (let Hred1 := fresh "Hred" in 
          (destruct (id_ok_is_R_xml_0_constructor H (refl_equal _)) as 
           [x12 [Heq Hred1]];
            (discriminate Heq)||
            (injection Heq;intros ;subst;clear Heq;clear H;
              impossible_star_reduction_R_xml_0 ))))
    | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
         _ (algebra.Alg.Term algebra.F.id_mark (?x12::nil)) |- _ =>
     let x12 := fresh "x" in 
      (let Heq := fresh "Heq" in 
        (let Hred1 := fresh "Hred" in 
          (destruct (id_mark_is_R_xml_0_constructor H (refl_equal _)) as 
           [x12 [Heq Hred1]];
            (discriminate Heq)||
            (injection Heq;intros ;subst;clear Heq;clear H;
              impossible_star_reduction_R_xml_0 ))))
    | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
         _ (algebra.Alg.Term algebra.F.id_u nil) |- _ =>
     let Heq := fresh "Heq" in 
      (set (Heq:=id_u_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_a nil) |- _ =>
     let Heq := fresh "Heq" in 
      (set (Heq:=id_a_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_o nil) |- _ =>
     let Heq := fresh "Heq" in 
      (set (Heq:=id_o_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_tt nil) |- _ =>
     let Heq := fresh "Heq" in 
      (set (Heq:=id_tt_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_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_e nil) |- _ =>
     let Heq := fresh "Heq" in 
      (set (Heq:=id_e_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_i nil) |- _ =>
     let Heq := fresh "Heq" in 
      (set (Heq:=id_i_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_ok (?x12::nil)) |- _ =>
     let x12 := fresh "x" in 
      (let Heq := fresh "Heq" in 
        (let Hred1 := fresh "Hred" in 
          (destruct (id_ok_is_R_xml_0_constructor H (refl_equal _)) as 
           [x12 [Heq Hred1]];
            (discriminate Heq)||
            (injection Heq;intros ;subst;clear Heq;clear H;
              try (simplify_star_reduction_R_xml_0 )))))
    | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
         _ (algebra.Alg.Term algebra.F.id_mark (?x12::nil)) |- _ =>
     let x12 := fresh "x" in 
      (let Heq := fresh "Heq" in 
        (let Hred1 := fresh "Hred" in 
          (destruct (id_mark_is_R_xml_0_constructor H (refl_equal _)) as 
           [x12 [Heq Hred1]];
            (discriminate Heq)||
            (injection Heq;intros ;subst;clear Heq;clear H;
              try (simplify_star_reduction_R_xml_0 )))))
    | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) 
         _ (algebra.Alg.Term algebra.F.id_u nil) |- _ =>
     let Heq := fresh "Heq" in 
      (set (Heq:=id_u_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_a nil) |- _ =>
     let Heq := fresh "Heq" in 
      (set (Heq:=id_a_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_o nil) |- _ =>
     let Heq := fresh "Heq" in 
      (set (Heq:=id_o_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_tt nil) |- _ =>
     let Heq := fresh "Heq" in 
      (set (Heq:=id_tt_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_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_e nil) |- _ =>
     let Heq := fresh "Heq" in 
      (set (Heq:=id_e_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_U71 : A ->A ->A.
   
   Hypothesis P_id_isList : A ->A.
   
   Hypothesis P_id_i : A.
   
   Hypothesis P_id_U11 : A ->A.
   
   Hypothesis P_id_isQid : A ->A.
   
   Hypothesis P_id_isNeList : A ->A.
   
   Hypothesis P_id_ok : A ->A.
   
   Hypothesis P_id_mark : A ->A.
   
   Hypothesis P_id_isPal : A ->A.
   
   Hypothesis P_id_U41 : A ->A ->A.
   
   Hypothesis P_id_u : A.
   
   Hypothesis P_id_U21 : A ->A ->A.
   
   Hypothesis P_id_a : A.
   
   Hypothesis P_id_U52 : A ->A.
   
   Hypothesis P_id___ : A ->A ->A.
   
   Hypothesis P_id_U72 : A ->A.
   
   Hypothesis P_id_U31 : A ->A.
   
   Hypothesis P_id_o : A.
   
   Hypothesis P_id_tt : A.
   
   Hypothesis P_id_isNePal : A ->A.
   
   Hypothesis P_id_U51 : A ->A ->A.
   
   Hypothesis P_id_top : A ->A.
   
   Hypothesis P_id_nil : A.
   
   Hypothesis P_id_U81 : A ->A.
   
   Hypothesis P_id_U42 : A ->A.
   
   Hypothesis P_id_proper : A ->A.
   
   Hypothesis P_id_U22 : A ->A.
   
   Hypothesis P_id_e : A.
   
   Hypothesis P_id_U61 : A ->A.
   
   Hypothesis P_id_active_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   
   Hypothesis P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (A0 <= x15)/\ (x15 <= x14) ->
      (A0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   
   Hypothesis P_id_isList_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   
   Hypothesis P_id_U11_monotonic :
    forall x12 x13, (A0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   
   Hypothesis P_id_isQid_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   
   Hypothesis P_id_isNeList_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   
   Hypothesis P_id_ok_monotonic :
    forall x12 x13, (A0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   
   Hypothesis P_id_mark_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   
   Hypothesis P_id_isPal_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   
   Hypothesis P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (A0 <= x15)/\ (x15 <= x14) ->
      (A0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   
   Hypothesis P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (A0 <= x15)/\ (x15 <= x14) ->
      (A0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   
   Hypothesis P_id_U52_monotonic :
    forall x12 x13, (A0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   
   Hypothesis P_id____monotonic :
    forall x12 x14 x13 x15, 
     (A0 <= x15)/\ (x15 <= x14) ->
      (A0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   
   Hypothesis P_id_U72_monotonic :
    forall x12 x13, (A0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   
   Hypothesis P_id_U31_monotonic :
    forall x12 x13, (A0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   
   Hypothesis P_id_isNePal_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   
   Hypothesis P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (A0 <= x15)/\ (x15 <= x14) ->
      (A0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   
   Hypothesis P_id_top_monotonic :
    forall x12 x13, (A0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   
   Hypothesis P_id_U81_monotonic :
    forall x12 x13, (A0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   
   Hypothesis P_id_U42_monotonic :
    forall x12 x13, (A0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   
   Hypothesis P_id_proper_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   
   Hypothesis P_id_U22_monotonic :
    forall x12 x13, (A0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   
   Hypothesis P_id_U61_monotonic :
    forall x12 x13, (A0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   
   Hypothesis P_id_active_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_active x12.
   
   Hypothesis P_id_U71_bounded :
    forall x12 x13, (A0 <= x12) ->(A0 <= x13) ->A0 <= P_id_U71 x13 x12.
   
   Hypothesis P_id_isList_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_isList x12.
   
   Hypothesis P_id_i_bounded : A0 <= P_id_i .
   
   Hypothesis P_id_U11_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_U11 x12.
   
   Hypothesis P_id_isQid_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_isQid x12.
   
   Hypothesis P_id_isNeList_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_isNeList x12.
   
   Hypothesis P_id_ok_bounded : forall x12, (A0 <= x12) ->A0 <= P_id_ok x12.
   
   Hypothesis P_id_mark_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_mark x12.
   
   Hypothesis P_id_isPal_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_isPal x12.
   
   Hypothesis P_id_U41_bounded :
    forall x12 x13, (A0 <= x12) ->(A0 <= x13) ->A0 <= P_id_U41 x13 x12.
   
   Hypothesis P_id_u_bounded : A0 <= P_id_u .
   
   Hypothesis P_id_U21_bounded :
    forall x12 x13, (A0 <= x12) ->(A0 <= x13) ->A0 <= P_id_U21 x13 x12.
   
   Hypothesis P_id_a_bounded : A0 <= P_id_a .
   
   Hypothesis P_id_U52_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_U52 x12.
   
   Hypothesis P_id____bounded :
    forall x12 x13, (A0 <= x12) ->(A0 <= x13) ->A0 <= P_id___ x13 x12.
   
   Hypothesis P_id_U72_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_U72 x12.
   
   Hypothesis P_id_U31_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_U31 x12.
   
   Hypothesis P_id_o_bounded : A0 <= P_id_o .
   
   Hypothesis P_id_tt_bounded : A0 <= P_id_tt .
   
   Hypothesis P_id_isNePal_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_isNePal x12.
   
   Hypothesis P_id_U51_bounded :
    forall x12 x13, (A0 <= x12) ->(A0 <= x13) ->A0 <= P_id_U51 x13 x12.
   
   Hypothesis P_id_top_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_top x12.
   
   Hypothesis P_id_nil_bounded : A0 <= P_id_nil .
   
   Hypothesis P_id_U81_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_U81 x12.
   
   Hypothesis P_id_U42_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_U42 x12.
   
   Hypothesis P_id_proper_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_proper x12.
   
   Hypothesis P_id_U22_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_U22 x12.
   
   Hypothesis P_id_e_bounded : A0 <= P_id_e .
   
   Hypothesis P_id_U61_bounded :
    forall x12, (A0 <= x12) ->A0 <= P_id_U61 x12.
   
   Fixpoint measure t { struct t }  := 
     match t with
       | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
        P_id_active (measure x12)
       | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
        P_id_U71 (measure x13) (measure x12)
       | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
        P_id_isList (measure x12)
       | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
       | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
        P_id_U11 (measure x12)
       | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
        P_id_isQid (measure x12)
       | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
        P_id_isNeList (measure x12)
       | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
        P_id_ok (measure x12)
       | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
        P_id_mark (measure x12)
       | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
        P_id_isPal (measure x12)
       | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
        P_id_U41 (measure x13) (measure x12)
       | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
       | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
        P_id_U21 (measure x13) (measure x12)
       | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
       | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
        P_id_U52 (measure x12)
       | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
        P_id___ (measure x13) (measure x12)
       | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
        P_id_U72 (measure x12)
       | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
        P_id_U31 (measure x12)
       | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
       | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
       | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
        P_id_isNePal (measure x12)
       | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
        P_id_U51 (measure x13) (measure x12)
       | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
        P_id_top (measure x12)
       | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
       | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
        P_id_U81 (measure x12)
       | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
        P_id_U42 (measure x12)
       | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
        P_id_proper (measure x12)
       | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
        P_id_U22 (measure x12)
       | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
       | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
        P_id_U61 (measure x12)
       | _ => A0
       end.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71 => P_id_U71
       | algebra.F.id_isList => P_id_isList
       | algebra.F.id_i => P_id_i
       | algebra.F.id_U11 => P_id_U11
       | algebra.F.id_isQid => P_id_isQid
       | algebra.F.id_isNeList => P_id_isNeList
       | algebra.F.id_ok => P_id_ok
       | algebra.F.id_mark => P_id_mark
       | algebra.F.id_isPal => P_id_isPal
       | algebra.F.id_U41 => P_id_U41
       | algebra.F.id_u => P_id_u
       | algebra.F.id_U21 => P_id_U21
       | algebra.F.id_a => P_id_a
       | algebra.F.id_U52 => P_id_U52
       | algebra.F.id___ => P_id___
       | algebra.F.id_U72 => P_id_U72
       | algebra.F.id_U31 => P_id_U31
       | algebra.F.id_o => P_id_o
       | algebra.F.id_tt => P_id_tt
       | algebra.F.id_isNePal => P_id_isNePal
       | algebra.F.id_U51 => P_id_U51
       | algebra.F.id_top => P_id_top
       | algebra.F.id_nil => P_id_nil
       | algebra.F.id_U81 => P_id_U81
       | algebra.F.id_U42 => P_id_U42
       | algebra.F.id_proper => P_id_proper
       | algebra.F.id_U22 => P_id_U22
       | algebra.F.id_e => P_id_e
       | algebra.F.id_U61 => P_id_U61
       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_U71 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_isList =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_i => match l with
                                    | nil => _
                                    | _::_ => _
                                    end
              | algebra.F.id_U11 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_isQid =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_isNeList =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_ok =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_mark =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_isPal =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_U41 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_u => match l with
                                    | nil => _
                                    | _::_ => _
                                    end
              | algebra.F.id_U21 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_a => match l with
                                    | nil => _
                                    | _::_ => _
                                    end
              | algebra.F.id_U52 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id___ =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_U72 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_U31 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_o => match l with
                                    | nil => _
                                    | _::_ => _
                                    end
              | algebra.F.id_tt => match l with
                                     | nil => _
                                     | _::_ => _
                                     end
              | algebra.F.id_isNePal =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_U51 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_top =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_nil => match l with
                                      | nil => _
                                      | _::_ => _
                                      end
              | algebra.F.id_U81 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_U42 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_proper =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_U22 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_e => match l with
                                    | nil => _
                                    | _::_ => _
                                    end
              | algebra.F.id_U61 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              end;simpl in |-*;unfold eq_rect_r, eq_rect, sym_eq in |-*;
      try (reflexivity );f_equal ;auto.
   Qed.
   
   Lemma measure_bounded : forall t, A0 <= measure t.
   Proof.
     intros t.
     rewrite same_measure in |-*.
     apply (InterpGen.measure_bounded Aop).
     intros f.
     case f.
     vm_compute in |-*;intros ;apply P_id_active_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U71_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_isList_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_i_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U11_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_isQid_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_isNeList_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_ok_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_mark_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_isPal_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U41_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U21_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_a_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U52_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id____bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U72_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U31_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_o_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_tt_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_isNePal_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U51_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_top_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_nil_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U81_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U42_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_proper_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U22_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_e_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U61_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_U71_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_isList_monotonic;assumption.
     vm_compute in |-*;apply (Aop.(le_refl)).
     vm_compute in |-*;intros ;apply P_id_U11_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_isQid_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_isNeList_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_ok_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_mark_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_isPal_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_U41_monotonic;assumption.
     vm_compute in |-*;apply (Aop.(le_refl)).
     vm_compute in |-*;intros ;apply P_id_U21_monotonic;assumption.
     vm_compute in |-*;apply (Aop.(le_refl)).
     vm_compute in |-*;intros ;apply P_id_U52_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id____monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_U72_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_U31_monotonic;assumption.
     vm_compute in |-*;apply (Aop.(le_refl)).
     vm_compute in |-*;apply (Aop.(le_refl)).
     vm_compute in |-*;intros ;apply P_id_isNePal_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_U51_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_top_monotonic;assumption.
     vm_compute in |-*;apply (Aop.(le_refl)).
     vm_compute in |-*;intros ;apply P_id_U81_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_U42_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_proper_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_U22_monotonic;assumption.
     vm_compute in |-*;apply (Aop.(le_refl)).
     vm_compute in |-*;intros ;apply P_id_U61_monotonic;assumption.
     intros f.
     case f.
     vm_compute in |-*;intros ;apply P_id_active_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U71_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_isList_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_i_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U11_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_isQid_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_isNeList_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_ok_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_mark_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_isPal_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U41_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U21_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_a_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U52_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id____bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U72_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U31_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_o_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_tt_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_isNePal_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U51_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_top_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_nil_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U81_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U42_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_proper_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U22_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_e_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U61_bounded;assumption.
     intros .
     do 2 (rewrite  <- same_measure in |-*).
     apply rules_monotonic;assumption.
     assumption.
   Qed.
   
   Hypothesis P_id_U22_hat_1 : A ->A.
   
   Hypothesis P_id_ISNEPAL : A ->A.
   
   Hypothesis P_id_U21_hat_1 : A ->A ->A.
   
   Hypothesis P_id_U52_hat_1 : A ->A.
   
   Hypothesis P_id_U51_hat_1 : A ->A ->A.
   
   Hypothesis P_id_U42_hat_1 : A ->A.
   
   Hypothesis P_id_TOP : A ->A.
   
   Hypothesis P_id_ISQID : A ->A.
   
   Hypothesis P_id_ISPAL : A ->A.
   
   Hypothesis P_id_U71_hat_1 : A ->A ->A.
   
   Hypothesis P_id_ACTIVE : A ->A.
   
   Hypothesis P_id_ISLIST : A ->A.
   
   Hypothesis P_id_PROPER : A ->A.
   
   Hypothesis P_id_U31_hat_1 : A ->A.
   
   Hypothesis P_id_U72_hat_1 : A ->A.
   
   Hypothesis P_id_U61_hat_1 : A ->A.
   
   Hypothesis P_id_ISNELIST : A ->A.
   
   Hypothesis P_id_U41_hat_1 : A ->A ->A.
   
   Hypothesis P_id_U11_hat_1 : A ->A.
   
   Hypothesis P_id_U81_hat_1 : A ->A.
   
   Hypothesis P_id____hat_1 : A ->A ->A.
   
   Hypothesis P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   
   Hypothesis P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   
   Hypothesis P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (A0 <= x15)/\ (x15 <= x14) ->
      (A0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   
   Hypothesis P_id_U52_hat_1_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   
   Hypothesis P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (A0 <= x15)/\ (x15 <= x14) ->
      (A0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   
   Hypothesis P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   
   Hypothesis P_id_TOP_monotonic :
    forall x12 x13, (A0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   
   Hypothesis P_id_ISQID_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   
   Hypothesis P_id_ISPAL_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   
   Hypothesis P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (A0 <= x15)/\ (x15 <= x14) ->
      (A0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   
   Hypothesis P_id_ACTIVE_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   
   Hypothesis P_id_ISLIST_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   
   Hypothesis P_id_PROPER_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   
   Hypothesis P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   
   Hypothesis P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   
   Hypothesis P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   
   Hypothesis P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   
   Hypothesis P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (A0 <= x15)/\ (x15 <= x14) ->
      (A0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   
   Hypothesis P_id_U11_hat_1_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   
   Hypothesis P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (A0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   
   Hypothesis P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (A0 <= x15)/\ (x15 <= x14) ->
      (A0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   
   Definition marked_measure t := 
     match t with
       | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
        P_id_U22_hat_1 (measure x12)
       | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
        P_id_ISNEPAL (measure x12)
       | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
        P_id_U21_hat_1 (measure x13) (measure x12)
       | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
        P_id_U52_hat_1 (measure x12)
       | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
        P_id_U51_hat_1 (measure x13) (measure x12)
       | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
        P_id_U42_hat_1 (measure x12)
       | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
        P_id_TOP (measure x12)
       | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
        P_id_ISQID (measure x12)
       | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
        P_id_ISPAL (measure x12)
       | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
        P_id_U71_hat_1 (measure x13) (measure x12)
       | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
        P_id_ACTIVE (measure x12)
       | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
        P_id_ISLIST (measure x12)
       | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
        P_id_PROPER (measure x12)
       | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
        P_id_U31_hat_1 (measure x12)
       | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
        P_id_U72_hat_1 (measure x12)
       | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
        P_id_U61_hat_1 (measure x12)
       | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
        P_id_ISNELIST (measure x12)
       | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
        P_id_U41_hat_1 (measure x13) (measure x12)
       | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
        P_id_U11_hat_1 (measure x12)
       | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
        P_id_U81_hat_1 (measure x12)
       | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
        P_id____hat_1 (measure x13) (measure x12)
       | _ => 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 x12;apply (P_id_U61_hat_1 x12).
     
     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 x12;apply (P_id_U22_hat_1 x12).
     
     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 x12;apply (P_id_PROPER x12).
     
     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 x12;apply (P_id_U42_hat_1 x12).
     
     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 x12;apply (P_id_U81_hat_1 x12).
     
     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 x12;apply (P_id_TOP x12).
     
     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 x13 x12;apply (P_id_U51_hat_1 x13 x12).
     
     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 x12;apply (P_id_ISNEPAL x12).
     
     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 x12;apply (P_id_U31_hat_1 x12).
     
     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 x12;apply (P_id_U72_hat_1 x12).
     
     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 x13 x12;apply (P_id____hat_1 x13 x12).
     
     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 x12;apply (P_id_U52_hat_1 x12).
     
     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 x13 x12;apply (P_id_U21_hat_1 x13 x12).
     
     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 x13 x12;apply (P_id_U41_hat_1 x13 x12).
     
     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 x12;apply (P_id_ISPAL x12).
     
     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 x12;apply (P_id_ISNELIST x12).
     
     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 x12;apply (P_id_ISQID x12).
     
     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 x12;apply (P_id_U11_hat_1 x12).
     
     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 x12;apply (P_id_ISLIST x12).
     
     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 x13 x12;apply (P_id_U71_hat_1 x13 x12).
     
     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 x12;apply (P_id_ACTIVE x12).
     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_U71 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_isList =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_i => match l with
                                    | nil => _
                                    | _::_ => _
                                    end
              | algebra.F.id_U11 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_isQid =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_isNeList =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_ok =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_mark =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_isPal =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_U41 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_u => match l with
                                    | nil => _
                                    | _::_ => _
                                    end
              | algebra.F.id_U21 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_a => match l with
                                    | nil => _
                                    | _::_ => _
                                    end
              | algebra.F.id_U52 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id___ =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_U72 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_U31 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_o => match l with
                                    | nil => _
                                    | _::_ => _
                                    end
              | algebra.F.id_tt => match l with
                                     | nil => _
                                     | _::_ => _
                                     end
              | algebra.F.id_isNePal =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_U51 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::nil => _
                 | _::_::_::_ => _
                 end
              | algebra.F.id_top =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_nil => match l with
                                      | nil => _
                                      | _::_ => _
                                      end
              | algebra.F.id_U81 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_U42 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_proper =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_U22 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              | algebra.F.id_e => match l with
                                    | nil => _
                                    | _::_ => _
                                    end
              | algebra.F.id_U61 =>
               match l with
                 | nil => _
                 | _::nil => _
                 | _::_::_ => _
                 end
              end.
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     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 .
     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 .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     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 .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     vm_compute in |-*;reflexivity .
     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 .
     vm_compute in |-*;reflexivity .
     
     lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ;
      apply same_measure.
     vm_compute in |-*;reflexivity .
   Qed.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_isList_monotonic;assumption.
     vm_compute in |-*;apply (Aop.(le_refl)).
     vm_compute in |-*;intros ;apply P_id_U11_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_isQid_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_isNeList_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_ok_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_mark_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_isPal_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_U41_monotonic;assumption.
     vm_compute in |-*;apply (Aop.(le_refl)).
     vm_compute in |-*;intros ;apply P_id_U21_monotonic;assumption.
     vm_compute in |-*;apply (Aop.(le_refl)).
     vm_compute in |-*;intros ;apply P_id_U52_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id____monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_U72_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_U31_monotonic;assumption.
     vm_compute in |-*;apply (Aop.(le_refl)).
     vm_compute in |-*;apply (Aop.(le_refl)).
     vm_compute in |-*;intros ;apply P_id_isNePal_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_U51_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_top_monotonic;assumption.
     vm_compute in |-*;apply (Aop.(le_refl)).
     vm_compute in |-*;intros ;apply P_id_U81_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_U42_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_proper_monotonic;assumption.
     vm_compute in |-*;intros ;apply P_id_U22_monotonic;assumption.
     vm_compute in |-*;apply (Aop.(le_refl)).
     vm_compute in |-*;intros ;apply P_id_U61_monotonic;assumption.
     clear f.
     intros f.
     case f.
     vm_compute in |-*;intros ;apply P_id_active_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U71_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_isList_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_i_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U11_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_isQid_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_isNeList_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_ok_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_mark_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_isPal_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U41_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_u_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U21_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_a_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U52_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id____bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U72_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U31_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_o_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_tt_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_isNePal_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U51_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_top_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_nil_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U81_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U42_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_proper_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U22_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_e_bounded;assumption.
     vm_compute in |-*;intros ;apply P_id_U61_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_U61_hat_1_monotonic;assumption.
     
     match type of H with 
       | orb ?a ?b = true => 
         assert (H':{a = true}+{b = true});[
           revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]|
             clear H;destruct H' as [H|H]]
     end .
     
     match type of H with 
       | _ _ ?t = true =>
         generalize (algebra.F.symb_eq_bool_ok f t);
           unfold algebra.Alg.eq_symb_bool in H;
           rewrite H;clear H;intros ;subst
     end .
     vm_compute in |-*;intros ;apply P_id_U22_hat_1_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_PROPER_monotonic;assumption.
     
     match type of H with 
       | orb ?a ?b = true => 
         assert (H':{a = true}+{b = true});[
           revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]|
             clear H;destruct H' as [H|H]]
     end .
     
     match type of H with 
       | _ _ ?t = true =>
         generalize (algebra.F.symb_eq_bool_ok f t);
           unfold algebra.Alg.eq_symb_bool in H;
           rewrite H;clear H;intros ;subst
     end .
     vm_compute in |-*;intros ;apply P_id_U42_hat_1_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_U81_hat_1_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_TOP_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_U51_hat_1_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_ISNEPAL_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_U31_hat_1_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_U72_hat_1_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____hat_1_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_U52_hat_1_monotonic;assumption.
     
     match type of H with 
       | orb ?a ?b = true => 
         assert (H':{a = true}+{b = true});[
           revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]|
             clear H;destruct H' as [H|H]]
     end .
     
     match type of H with 
       | _ _ ?t = true =>
         generalize (algebra.F.symb_eq_bool_ok f t);
           unfold algebra.Alg.eq_symb_bool in H;
           rewrite H;clear H;intros ;subst
     end .
     vm_compute in |-*;intros ;apply P_id_U21_hat_1_monotonic;assumption.
     
     match type of H with 
       | orb ?a ?b = true => 
         assert (H':{a = true}+{b = true});[
           revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]|
             clear H;destruct H' as [H|H]]
     end .
     
     match type of H with 
       | _ _ ?t = true =>
         generalize (algebra.F.symb_eq_bool_ok f t);
           unfold algebra.Alg.eq_symb_bool in H;
           rewrite H;clear H;intros ;subst
     end .
     vm_compute in |-*;intros ;apply P_id_U41_hat_1_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_ISPAL_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_ISNELIST_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_ISQID_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_U11_hat_1_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_ISLIST_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_U71_hat_1_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_U71 : Z ->Z ->Z.
   
   Hypothesis P_id_isList : Z ->Z.
   
   Hypothesis P_id_i : Z.
   
   Hypothesis P_id_U11 : Z ->Z.
   
   Hypothesis P_id_isQid : Z ->Z.
   
   Hypothesis P_id_isNeList : Z ->Z.
   
   Hypothesis P_id_ok : Z ->Z.
   
   Hypothesis P_id_mark : Z ->Z.
   
   Hypothesis P_id_isPal : Z ->Z.
   
   Hypothesis P_id_U41 : Z ->Z ->Z.
   
   Hypothesis P_id_u : Z.
   
   Hypothesis P_id_U21 : Z ->Z ->Z.
   
   Hypothesis P_id_a : Z.
   
   Hypothesis P_id_U52 : Z ->Z.
   
   Hypothesis P_id___ : Z ->Z ->Z.
   
   Hypothesis P_id_U72 : Z ->Z.
   
   Hypothesis P_id_U31 : Z ->Z.
   
   Hypothesis P_id_o : Z.
   
   Hypothesis P_id_tt : Z.
   
   Hypothesis P_id_isNePal : Z ->Z.
   
   Hypothesis P_id_U51 : Z ->Z ->Z.
   
   Hypothesis P_id_top : Z ->Z.
   
   Hypothesis P_id_nil : Z.
   
   Hypothesis P_id_U81 : Z ->Z.
   
   Hypothesis P_id_U42 : Z ->Z.
   
   Hypothesis P_id_proper : Z ->Z.
   
   Hypothesis P_id_U22 : Z ->Z.
   
   Hypothesis P_id_e : Z.
   
   Hypothesis P_id_U61 : Z ->Z.
   
   Hypothesis P_id_active_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   
   Hypothesis P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (min_value <= x15)/\ (x15 <= x14) ->
      (min_value <= x13)/\ (x13 <= x12) ->
       P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   
   Hypothesis P_id_isList_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   
   Hypothesis P_id_U11_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   
   Hypothesis P_id_isQid_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   
   Hypothesis P_id_isNeList_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->
      P_id_isNeList x13 <= P_id_isNeList x12.
   
   Hypothesis P_id_ok_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   
   Hypothesis P_id_mark_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   
   Hypothesis P_id_isPal_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   
   Hypothesis P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (min_value <= x15)/\ (x15 <= x14) ->
      (min_value <= x13)/\ (x13 <= x12) ->
       P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   
   Hypothesis P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (min_value <= x15)/\ (x15 <= x14) ->
      (min_value <= x13)/\ (x13 <= x12) ->
       P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   
   Hypothesis P_id_U52_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   
   Hypothesis P_id____monotonic :
    forall x12 x14 x13 x15, 
     (min_value <= x15)/\ (x15 <= x14) ->
      (min_value <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   
   Hypothesis P_id_U72_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   
   Hypothesis P_id_U31_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   
   Hypothesis P_id_isNePal_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   
   Hypothesis P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (min_value <= x15)/\ (x15 <= x14) ->
      (min_value <= x13)/\ (x13 <= x12) ->
       P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   
   Hypothesis P_id_top_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   
   Hypothesis P_id_U81_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   
   Hypothesis P_id_U42_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   
   Hypothesis P_id_proper_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   
   Hypothesis P_id_U22_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   
   Hypothesis P_id_U61_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   
   Hypothesis P_id_active_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_active x12.
   
   Hypothesis P_id_U71_bounded :
    forall x12 x13, 
     (min_value <= x12) ->(min_value <= x13) ->min_value <= P_id_U71 x13 x12.
   
   Hypothesis P_id_isList_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_isList x12.
   
   Hypothesis P_id_i_bounded : min_value <= P_id_i .
   
   Hypothesis P_id_U11_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_U11 x12.
   
   Hypothesis P_id_isQid_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_isQid x12.
   
   Hypothesis P_id_isNeList_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_isNeList x12.
   
   Hypothesis P_id_ok_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_ok x12.
   
   Hypothesis P_id_mark_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_mark x12.
   
   Hypothesis P_id_isPal_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_isPal x12.
   
   Hypothesis P_id_U41_bounded :
    forall x12 x13, 
     (min_value <= x12) ->(min_value <= x13) ->min_value <= P_id_U41 x13 x12.
   
   Hypothesis P_id_u_bounded : min_value <= P_id_u .
   
   Hypothesis P_id_U21_bounded :
    forall x12 x13, 
     (min_value <= x12) ->(min_value <= x13) ->min_value <= P_id_U21 x13 x12.
   
   Hypothesis P_id_a_bounded : min_value <= P_id_a .
   
   Hypothesis P_id_U52_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_U52 x12.
   
   Hypothesis P_id____bounded :
    forall x12 x13, 
     (min_value <= x12) ->(min_value <= x13) ->min_value <= P_id___ x13 x12.
   
   Hypothesis P_id_U72_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_U72 x12.
   
   Hypothesis P_id_U31_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_U31 x12.
   
   Hypothesis P_id_o_bounded : min_value <= P_id_o .
   
   Hypothesis P_id_tt_bounded : min_value <= P_id_tt .
   
   Hypothesis P_id_isNePal_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_isNePal x12.
   
   Hypothesis P_id_U51_bounded :
    forall x12 x13, 
     (min_value <= x12) ->(min_value <= x13) ->min_value <= P_id_U51 x13 x12.
   
   Hypothesis P_id_top_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_top x12.
   
   Hypothesis P_id_nil_bounded : min_value <= P_id_nil .
   
   Hypothesis P_id_U81_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_U81 x12.
   
   Hypothesis P_id_U42_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_U42 x12.
   
   Hypothesis P_id_proper_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_proper x12.
   
   Hypothesis P_id_U22_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_U22 x12.
   
   Hypothesis P_id_e_bounded : min_value <= P_id_e .
   
   Hypothesis P_id_U61_bounded :
    forall x12, (min_value <= x12) ->min_value <= P_id_U61 x12.
   
   Definition measure  := 
     Interp.measure min_value P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Hypothesis P_id_U22_hat_1 : Z ->Z.
   
   Hypothesis P_id_ISNEPAL : Z ->Z.
   
   Hypothesis P_id_U21_hat_1 : Z ->Z ->Z.
   
   Hypothesis P_id_U52_hat_1 : Z ->Z.
   
   Hypothesis P_id_U51_hat_1 : Z ->Z ->Z.
   
   Hypothesis P_id_U42_hat_1 : Z ->Z.
   
   Hypothesis P_id_TOP : Z ->Z.
   
   Hypothesis P_id_ISQID : Z ->Z.
   
   Hypothesis P_id_ISPAL : Z ->Z.
   
   Hypothesis P_id_U71_hat_1 : Z ->Z ->Z.
   
   Hypothesis P_id_ACTIVE : Z ->Z.
   
   Hypothesis P_id_ISLIST : Z ->Z.
   
   Hypothesis P_id_PROPER : Z ->Z.
   
   Hypothesis P_id_U31_hat_1 : Z ->Z.
   
   Hypothesis P_id_U72_hat_1 : Z ->Z.
   
   Hypothesis P_id_U61_hat_1 : Z ->Z.
   
   Hypothesis P_id_ISNELIST : Z ->Z.
   
   Hypothesis P_id_U41_hat_1 : Z ->Z ->Z.
   
   Hypothesis P_id_U11_hat_1 : Z ->Z.
   
   Hypothesis P_id_U81_hat_1 : Z ->Z.
   
   Hypothesis P_id____hat_1 : Z ->Z ->Z.
   
   Hypothesis P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->
      P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   
   Hypothesis P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   
   Hypothesis P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (min_value <= x15)/\ (x15 <= x14) ->
      (min_value <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   
   Hypothesis P_id_U52_hat_1_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->
      P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   
   Hypothesis P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (min_value <= x15)/\ (x15 <= x14) ->
      (min_value <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   
   Hypothesis P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->
      P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   
   Hypothesis P_id_TOP_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   
   Hypothesis P_id_ISQID_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   
   Hypothesis P_id_ISPAL_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   
   Hypothesis P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (min_value <= x15)/\ (x15 <= x14) ->
      (min_value <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   
   Hypothesis P_id_ACTIVE_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   
   Hypothesis P_id_ISLIST_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   
   Hypothesis P_id_PROPER_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   
   Hypothesis P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->
      P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   
   Hypothesis P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->
      P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   
   Hypothesis P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->
      P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   
   Hypothesis P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->
      P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   
   Hypothesis P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (min_value <= x15)/\ (x15 <= x14) ->
      (min_value <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   
   Hypothesis P_id_U11_hat_1_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->
      P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   
   Hypothesis P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (min_value <= x13)/\ (x13 <= x12) ->
      P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   
   Hypothesis P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (min_value <= x15)/\ (x15 <= x14) ->
      (min_value <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   
   Definition marked_measure  := 
     Interp.marked_measure min_value P_id_active P_id_U71 P_id_isList 
      P_id_i P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   End S.
End InterpZ.

Module WF_R_xml_0_deep_rew.
 Inductive DP_R_xml_0  :
  algebra.Alg.term ->algebra.Alg.term ->Prop := 
    (* <active(__(__(X_,Y_),Z_)),__(X_,__(Y_,Z_))> *)
   | DP_R_xml_0_0 :
    forall x12 x2 x1 x3, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Term algebra.F.id___ 
        (x1::x2::nil))::x3::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id___ (x1::(algebra.Alg.Term 
                  algebra.F.id___ (x2::x3::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(__(__(X_,Y_),Z_)),__(Y_,Z_)> *)
   | DP_R_xml_0_1 :
    forall x12 x2 x1 x3, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Term algebra.F.id___ 
        (x1::x2::nil))::x3::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id___ (x2::x3::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U21(tt,V2_)),U22(isList(V2_))> *)
   | DP_R_xml_0_2 :
    forall x4 x12, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Term algebra.F.id_tt 
        nil)::x4::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Term 
                  algebra.F.id_isList (x4::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U21(tt,V2_)),isList(V2_)> *)
   | DP_R_xml_0_3 :
    forall x4 x12, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Term algebra.F.id_tt 
        nil)::x4::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isList (x4::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U41(tt,V2_)),U42(isNeList(V2_))> *)
   | DP_R_xml_0_4 :
    forall x4 x12, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Term algebra.F.id_tt 
        nil)::x4::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Term 
                  algebra.F.id_isNeList (x4::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U41(tt,V2_)),isNeList(V2_)> *)
   | DP_R_xml_0_5 :
    forall x4 x12, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Term algebra.F.id_tt 
        nil)::x4::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isNeList (x4::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U51(tt,V2_)),U52(isList(V2_))> *)
   | DP_R_xml_0_6 :
    forall x4 x12, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Term algebra.F.id_tt 
        nil)::x4::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Term 
                  algebra.F.id_isList (x4::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U51(tt,V2_)),isList(V2_)> *)
   | DP_R_xml_0_7 :
    forall x4 x12, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Term algebra.F.id_tt 
        nil)::x4::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isList (x4::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U71(tt,P_)),U72(isPal(P_))> *)
   | DP_R_xml_0_8 :
    forall x12 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Term algebra.F.id_tt 
        nil)::x5::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U72 ((algebra.Alg.Term 
                  algebra.F.id_isPal (x5::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U71(tt,P_)),isPal(P_)> *)
   | DP_R_xml_0_9 :
    forall x12 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Term algebra.F.id_tt 
        nil)::x5::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isPal (x5::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(isList(V_)),U11(isNeList(V_))> *)
   | DP_R_xml_0_10 :
    forall x12 x6, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isList (x6::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Term 
                  algebra.F.id_isNeList (x6::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(isList(V_)),isNeList(V_)> *)
   | DP_R_xml_0_11 :
    forall x12 x6, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isList (x6::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isNeList (x6::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(isList(__(V1_,V2_))),U21(isList(V1_),V2_)> *)
   | DP_R_xml_0_12 :
    forall x4 x12 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isList ((algebra.Alg.Term 
        algebra.F.id___ (x7::x4::nil))::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Term 
                  algebra.F.id_isList (x7::nil))::x4::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(isList(__(V1_,V2_))),isList(V1_)> *)
   | DP_R_xml_0_13 :
    forall x4 x12 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isList ((algebra.Alg.Term 
        algebra.F.id___ (x7::x4::nil))::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isList (x7::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(isNeList(V_)),U31(isQid(V_))> *)
   | DP_R_xml_0_14 :
    forall x12 x6, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isNeList (x6::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Term 
                  algebra.F.id_isQid (x6::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(isNeList(V_)),isQid(V_)> *)
   | DP_R_xml_0_15 :
    forall x12 x6, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isNeList (x6::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isQid (x6::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
   
    (* <active(isNeList(__(V1_,V2_))),U41(isList(V1_),V2_)> *)
   | DP_R_xml_0_16 :
    forall x4 x12 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Term 
        algebra.F.id___ (x7::x4::nil))::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Term 
                  algebra.F.id_isList (x7::nil))::x4::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(isNeList(__(V1_,V2_))),isList(V1_)> *)
   | DP_R_xml_0_17 :
    forall x4 x12 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Term 
        algebra.F.id___ (x7::x4::nil))::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isList (x7::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
   
    (* <active(isNeList(__(V1_,V2_))),U51(isNeList(V1_),V2_)> *)
   | DP_R_xml_0_18 :
    forall x4 x12 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Term 
        algebra.F.id___ (x7::x4::nil))::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Term 
                  algebra.F.id_isNeList (x7::nil))::x4::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(isNeList(__(V1_,V2_))),isNeList(V1_)> *)
   | DP_R_xml_0_19 :
    forall x4 x12 x7, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Term 
        algebra.F.id___ (x7::x4::nil))::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isNeList (x7::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(isNePal(V_)),U61(isQid(V_))> *)
   | DP_R_xml_0_20 :
    forall x12 x6, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isNePal (x6::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Term 
                  algebra.F.id_isQid (x6::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(isNePal(V_)),isQid(V_)> *)
   | DP_R_xml_0_21 :
    forall x12 x6, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isNePal (x6::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isQid (x6::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
   
    (* <active(isNePal(__(I_,__(P_,I_)))),U71(isQid(I_),P_)> *)
   | DP_R_xml_0_22 :
    forall x8 x12 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isNePal ((algebra.Alg.Term 
        algebra.F.id___ (x8::(algebra.Alg.Term algebra.F.id___ (x5::
        x8::nil))::nil))::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Term 
                  algebra.F.id_isQid (x8::nil))::x5::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(isNePal(__(I_,__(P_,I_)))),isQid(I_)> *)
   | DP_R_xml_0_23 :
    forall x8 x12 x5, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isNePal ((algebra.Alg.Term 
        algebra.F.id___ (x8::(algebra.Alg.Term algebra.F.id___ (x5::
        x8::nil))::nil))::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isQid (x8::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(isPal(V_)),U81(isNePal(V_))> *)
   | DP_R_xml_0_24 :
    forall x12 x6, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isPal (x6::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U81 ((algebra.Alg.Term 
                  algebra.F.id_isNePal (x6::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(isPal(V_)),isNePal(V_)> *)
   | DP_R_xml_0_25 :
    forall x12 x6, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isPal (x6::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isNePal (x6::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(__(X1_,X2_)),__(active(X1_),X2_)> *)
   | DP_R_xml_0_26 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Term 
                  algebra.F.id_active (x9::nil))::x10::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(__(X1_,X2_)),active(X1_)> *)
   | DP_R_xml_0_27 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active (x9::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(__(X1_,X2_)),__(X1_,active(X2_))> *)
   | DP_R_xml_0_28 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id___ (x9::(algebra.Alg.Term 
                  algebra.F.id_active (x10::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(__(X1_,X2_)),active(X2_)> *)
   | DP_R_xml_0_29 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active (x10::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U11(X_)),U11(active(X_))> *)
   | DP_R_xml_0_30 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U11 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Term 
                  algebra.F.id_active (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U11(X_)),active(X_)> *)
   | DP_R_xml_0_31 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U11 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U21(X1_,X2_)),U21(active(X1_),X2_)> *)
   | DP_R_xml_0_32 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Term 
                  algebra.F.id_active (x9::nil))::x10::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U21(X1_,X2_)),active(X1_)> *)
   | DP_R_xml_0_33 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active (x9::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U22(X_)),U22(active(X_))> *)
   | DP_R_xml_0_34 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U22 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Term 
                  algebra.F.id_active (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U22(X_)),active(X_)> *)
   | DP_R_xml_0_35 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U22 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U31(X_)),U31(active(X_))> *)
   | DP_R_xml_0_36 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U31 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Term 
                  algebra.F.id_active (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U31(X_)),active(X_)> *)
   | DP_R_xml_0_37 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U31 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U41(X1_,X2_)),U41(active(X1_),X2_)> *)
   | DP_R_xml_0_38 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U41 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Term 
                  algebra.F.id_active (x9::nil))::x10::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U41(X1_,X2_)),active(X1_)> *)
   | DP_R_xml_0_39 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U41 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active (x9::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U42(X_)),U42(active(X_))> *)
   | DP_R_xml_0_40 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U42 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Term 
                  algebra.F.id_active (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U42(X_)),active(X_)> *)
   | DP_R_xml_0_41 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U42 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U51(X1_,X2_)),U51(active(X1_),X2_)> *)
   | DP_R_xml_0_42 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U51 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Term 
                  algebra.F.id_active (x9::nil))::x10::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U51(X1_,X2_)),active(X1_)> *)
   | DP_R_xml_0_43 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U51 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active (x9::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U52(X_)),U52(active(X_))> *)
   | DP_R_xml_0_44 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U52 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Term 
                  algebra.F.id_active (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U52(X_)),active(X_)> *)
   | DP_R_xml_0_45 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U52 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U61(X_)),U61(active(X_))> *)
   | DP_R_xml_0_46 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U61 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Term 
                  algebra.F.id_active (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U61(X_)),active(X_)> *)
   | DP_R_xml_0_47 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U61 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U71(X1_,X2_)),U71(active(X1_),X2_)> *)
   | DP_R_xml_0_48 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Term 
                  algebra.F.id_active (x9::nil))::x10::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U71(X1_,X2_)),active(X1_)> *)
   | DP_R_xml_0_49 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active (x9::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U72(X_)),U72(active(X_))> *)
   | DP_R_xml_0_50 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U72 ((algebra.Alg.Term 
                  algebra.F.id_active (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U72(X_)),active(X_)> *)
   | DP_R_xml_0_51 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U81(X_)),U81(active(X_))> *)
   | DP_R_xml_0_52 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U81 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U81 ((algebra.Alg.Term 
                  algebra.F.id_active (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <active(U81(X_)),active(X_)> *)
   | DP_R_xml_0_53 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U81 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_active (x12::nil))
    (* <__(mark(X1_),X2_),__(X1_,X2_)> *)
   | DP_R_xml_0_54 :
    forall x12 x10 x9 x13, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x9::nil)) x13) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x10 x12) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) 
        (algebra.Alg.Term algebra.F.id___ (x13::x12::nil))
    (* <__(X1_,mark(X2_)),__(X1_,X2_)> *)
   | DP_R_xml_0_55 :
    forall x12 x10 x9 x13, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                x9 x13) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x10::nil)) x12) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) 
        (algebra.Alg.Term algebra.F.id___ (x13::x12::nil))
    (* <U11(mark(X_)),U11(X_)> *)
   | DP_R_xml_0_56 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U11 (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_U11 (x12::nil))
    (* <U21(mark(X1_),X2_),U21(X1_,X2_)> *)
   | DP_R_xml_0_57 :
    forall x12 x10 x9 x13, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x9::nil)) x13) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x10 x12) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) 
        (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil))
    (* <U22(mark(X_)),U22(X_)> *)
   | DP_R_xml_0_58 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U22 (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_U22 (x12::nil))
    (* <U31(mark(X_)),U31(X_)> *)
   | DP_R_xml_0_59 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U31 (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_U31 (x12::nil))
    (* <U41(mark(X1_),X2_),U41(X1_,X2_)> *)
   | DP_R_xml_0_60 :
    forall x12 x10 x9 x13, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x9::nil)) x13) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x10 x12) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U41 (x9::x10::nil)) 
        (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil))
    (* <U42(mark(X_)),U42(X_)> *)
   | DP_R_xml_0_61 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U42 (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_U42 (x12::nil))
    (* <U51(mark(X1_),X2_),U51(X1_,X2_)> *)
   | DP_R_xml_0_62 :
    forall x12 x10 x9 x13, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x9::nil)) x13) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x10 x12) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U51 (x9::x10::nil)) 
        (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil))
    (* <U52(mark(X_)),U52(X_)> *)
   | DP_R_xml_0_63 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U52 (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_U52 (x12::nil))
    (* <U61(mark(X_)),U61(X_)> *)
   | DP_R_xml_0_64 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U61 (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_U61 (x12::nil))
    (* <U71(mark(X1_),X2_),U71(X1_,X2_)> *)
   | DP_R_xml_0_65 :
    forall x12 x10 x9 x13, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x9::nil)) x13) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x10 x12) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) 
        (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil))
    (* <U72(mark(X_)),U72(X_)> *)
   | DP_R_xml_0_66 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_U72 (x12::nil))
    (* <U81(mark(X_)),U81(X_)> *)
   | DP_R_xml_0_67 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U81 (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_U81 (x12::nil))
    (* <proper(__(X1_,X2_)),__(proper(X1_),proper(X2_))> *)
   | DP_R_xml_0_68 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Term 
                  algebra.F.id_proper (x9::nil))::(algebra.Alg.Term 
                  algebra.F.id_proper (x10::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(__(X1_,X2_)),proper(X1_)> *)
   | DP_R_xml_0_69 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x9::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(__(X1_,X2_)),proper(X2_)> *)
   | DP_R_xml_0_70 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x10::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U11(X_)),U11(proper(X_))> *)
   | DP_R_xml_0_71 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U11 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Term 
                  algebra.F.id_proper (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U11(X_)),proper(X_)> *)
   | DP_R_xml_0_72 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U11 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
   
    (* <proper(U21(X1_,X2_)),U21(proper(X1_),proper(X2_))> *)
   | DP_R_xml_0_73 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Term 
                  algebra.F.id_proper (x9::nil))::(algebra.Alg.Term 
                  algebra.F.id_proper (x10::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U21(X1_,X2_)),proper(X1_)> *)
   | DP_R_xml_0_74 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x9::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U21(X1_,X2_)),proper(X2_)> *)
   | DP_R_xml_0_75 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x10::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U22(X_)),U22(proper(X_))> *)
   | DP_R_xml_0_76 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U22 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Term 
                  algebra.F.id_proper (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U22(X_)),proper(X_)> *)
   | DP_R_xml_0_77 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U22 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(isList(X_)),isList(proper(X_))> *)
   | DP_R_xml_0_78 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isList ((algebra.Alg.Term 
                  algebra.F.id_proper (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(isList(X_)),proper(X_)> *)
   | DP_R_xml_0_79 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U31(X_)),U31(proper(X_))> *)
   | DP_R_xml_0_80 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U31 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Term 
                  algebra.F.id_proper (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U31(X_)),proper(X_)> *)
   | DP_R_xml_0_81 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U31 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
   
    (* <proper(U41(X1_,X2_)),U41(proper(X1_),proper(X2_))> *)
   | DP_R_xml_0_82 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U41 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Term 
                  algebra.F.id_proper (x9::nil))::(algebra.Alg.Term 
                  algebra.F.id_proper (x10::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U41(X1_,X2_)),proper(X1_)> *)
   | DP_R_xml_0_83 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U41 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x9::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U41(X1_,X2_)),proper(X2_)> *)
   | DP_R_xml_0_84 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U41 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x10::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U42(X_)),U42(proper(X_))> *)
   | DP_R_xml_0_85 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U42 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Term 
                  algebra.F.id_proper (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U42(X_)),proper(X_)> *)
   | DP_R_xml_0_86 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U42 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(isNeList(X_)),isNeList(proper(X_))> *)
   | DP_R_xml_0_87 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Term 
                  algebra.F.id_proper (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(isNeList(X_)),proper(X_)> *)
   | DP_R_xml_0_88 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
   
    (* <proper(U51(X1_,X2_)),U51(proper(X1_),proper(X2_))> *)
   | DP_R_xml_0_89 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U51 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Term 
                  algebra.F.id_proper (x9::nil))::(algebra.Alg.Term 
                  algebra.F.id_proper (x10::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U51(X1_,X2_)),proper(X1_)> *)
   | DP_R_xml_0_90 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U51 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x9::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U51(X1_,X2_)),proper(X2_)> *)
   | DP_R_xml_0_91 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U51 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x10::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U52(X_)),U52(proper(X_))> *)
   | DP_R_xml_0_92 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U52 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Term 
                  algebra.F.id_proper (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U52(X_)),proper(X_)> *)
   | DP_R_xml_0_93 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U52 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U61(X_)),U61(proper(X_))> *)
   | DP_R_xml_0_94 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U61 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Term 
                  algebra.F.id_proper (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U61(X_)),proper(X_)> *)
   | DP_R_xml_0_95 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U61 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
   
    (* <proper(U71(X1_,X2_)),U71(proper(X1_),proper(X2_))> *)
   | DP_R_xml_0_96 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Term 
                  algebra.F.id_proper (x9::nil))::(algebra.Alg.Term 
                  algebra.F.id_proper (x10::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U71(X1_,X2_)),proper(X1_)> *)
   | DP_R_xml_0_97 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x9::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U71(X1_,X2_)),proper(X2_)> *)
   | DP_R_xml_0_98 :
    forall x12 x10 x9, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x10::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U72(X_)),U72(proper(X_))> *)
   | DP_R_xml_0_99 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U72 ((algebra.Alg.Term 
                  algebra.F.id_proper (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U72(X_)),proper(X_)> *)
   | DP_R_xml_0_100 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(isPal(X_)),isPal(proper(X_))> *)
   | DP_R_xml_0_101 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isPal (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isPal ((algebra.Alg.Term 
                  algebra.F.id_proper (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(isPal(X_)),proper(X_)> *)
   | DP_R_xml_0_102 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isPal (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U81(X_)),U81(proper(X_))> *)
   | DP_R_xml_0_103 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U81 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U81 ((algebra.Alg.Term 
                  algebra.F.id_proper (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(U81(X_)),proper(X_)> *)
   | DP_R_xml_0_104 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_U81 (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(isQid(X_)),isQid(proper(X_))> *)
   | DP_R_xml_0_105 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isQid (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isQid ((algebra.Alg.Term 
                  algebra.F.id_proper (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(isQid(X_)),proper(X_)> *)
   | DP_R_xml_0_106 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isQid (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(isNePal(X_)),isNePal(proper(X_))> *)
   | DP_R_xml_0_107 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isNePal (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isNePal ((algebra.Alg.Term 
                  algebra.F.id_proper (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <proper(isNePal(X_)),proper(X_)> *)
   | DP_R_xml_0_108 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_isNePal (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    (* <__(ok(X1_),ok(X2_)),__(X1_,X2_)> *)
   | DP_R_xml_0_109 :
    forall x12 x10 x9 x13, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x9::nil)) 
       x13) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x10::nil)) x12) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) 
        (algebra.Alg.Term algebra.F.id___ (x13::x12::nil))
    (* <U11(ok(X_)),U11(X_)> *)
   | DP_R_xml_0_110 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x1::nil)) 
       x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U11 (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_U11 (x12::nil))
    (* <U21(ok(X1_),ok(X2_)),U21(X1_,X2_)> *)
   | DP_R_xml_0_111 :
    forall x12 x10 x9 x13, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x9::nil)) 
       x13) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x10::nil)) x12) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) 
        (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil))
    (* <U22(ok(X_)),U22(X_)> *)
   | DP_R_xml_0_112 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x1::nil)) 
       x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U22 (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_U22 (x12::nil))
    (* <isList(ok(X_)),isList(X_)> *)
   | DP_R_xml_0_113 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x1::nil)) 
       x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isList (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_isList (x12::nil))
    (* <U31(ok(X_)),U31(X_)> *)
   | DP_R_xml_0_114 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x1::nil)) 
       x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U31 (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_U31 (x12::nil))
    (* <U41(ok(X1_),ok(X2_)),U41(X1_,X2_)> *)
   | DP_R_xml_0_115 :
    forall x12 x10 x9 x13, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x9::nil)) 
       x13) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x10::nil)) x12) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U41 (x9::x10::nil)) 
        (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil))
    (* <U42(ok(X_)),U42(X_)> *)
   | DP_R_xml_0_116 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x1::nil)) 
       x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U42 (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_U42 (x12::nil))
    (* <isNeList(ok(X_)),isNeList(X_)> *)
   | DP_R_xml_0_117 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x1::nil)) 
       x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_isNeList (x12::nil))
    (* <U51(ok(X1_),ok(X2_)),U51(X1_,X2_)> *)
   | DP_R_xml_0_118 :
    forall x12 x10 x9 x13, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x9::nil)) 
       x13) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x10::nil)) x12) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U51 (x9::x10::nil)) 
        (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil))
    (* <U52(ok(X_)),U52(X_)> *)
   | DP_R_xml_0_119 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x1::nil)) 
       x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U52 (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_U52 (x12::nil))
    (* <U61(ok(X_)),U61(X_)> *)
   | DP_R_xml_0_120 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x1::nil)) 
       x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U61 (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_U61 (x12::nil))
    (* <U71(ok(X1_),ok(X2_)),U71(X1_,X2_)> *)
   | DP_R_xml_0_121 :
    forall x12 x10 x9 x13, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x9::nil)) 
       x13) ->
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x10::nil)) x12) ->
       DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) 
        (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil))
    (* <U72(ok(X_)),U72(X_)> *)
   | DP_R_xml_0_122 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x1::nil)) 
       x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_U72 (x12::nil))
    (* <isPal(ok(X_)),isPal(X_)> *)
   | DP_R_xml_0_123 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x1::nil)) 
       x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isPal (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_isPal (x12::nil))
    (* <U81(ok(X_)),U81(X_)> *)
   | DP_R_xml_0_124 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x1::nil)) 
       x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_U81 (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_U81 (x12::nil))
    (* <isQid(ok(X_)),isQid(X_)> *)
   | DP_R_xml_0_125 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x1::nil)) 
       x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isQid (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_isQid (x12::nil))
    (* <isNePal(ok(X_)),isNePal(X_)> *)
   | DP_R_xml_0_126 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x1::nil)) 
       x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_isNePal (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_isNePal (x12::nil))
    (* <top(mark(X_)),top(proper(X_))> *)
   | DP_R_xml_0_127 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_top ((algebra.Alg.Term 
                  algebra.F.id_proper (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_top (x12::nil))
    (* <top(mark(X_)),proper(X_)> *)
   | DP_R_xml_0_128 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                
       (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_top (x12::nil))
    (* <top(ok(X_)),top(active(X_))> *)
   | DP_R_xml_0_129 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x1::nil)) 
       x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_top ((algebra.Alg.Term 
                  algebra.F.id_active (x1::nil))::nil)) 
       (algebra.Alg.Term algebra.F.id_top (x12::nil))
    (* <top(ok(X_)),active(X_)> *)
   | DP_R_xml_0_130 :
    forall x12 x1, 
     (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                (algebra.Alg.Term algebra.F.id_ok (x1::nil)) 
       x12) ->
      DP_R_xml_0 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
       (algebra.Alg.Term algebra.F.id_top (x12::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 := 
     (* <isNePal(ok(X_)),isNePal(X_)> *)
    | DP_R_xml_0_scc_1_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
       DP_R_xml_0_scc_1 (algebra.Alg.Term algebra.F.id_isNePal (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_isNePal (x12::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 (x12:Z) := 2* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_isList (x12:Z) := 2 + 2* x12.
   
   Definition P_id_i  := 2.
   
   Definition P_id_U11 (x12:Z) := 3 + 1* x12.
   
   Definition P_id_isQid (x12:Z) := 2 + 2* x12.
   
   Definition P_id_isNeList (x12:Z) := 2* x12.
   
   Definition P_id_ok (x12:Z) := 1 + 1* x12.
   
   Definition P_id_mark (x12:Z) := 0.
   
   Definition P_id_isPal (x12:Z) := 1* x12.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_u  := 1.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_a  := 1.
   
   Definition P_id_U52 (x12:Z) := 2* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_U72 (x12:Z) := 1* x12.
   
   Definition P_id_U31 (x12:Z) := 2 + 2* x12.
   
   Definition P_id_o  := 1.
   
   Definition P_id_tt  := 2.
   
   Definition P_id_isNePal (x12:Z) := 2* x12.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 2.
   
   Definition P_id_U81 (x12:Z) := 2* x12.
   
   Definition P_id_U42 (x12:Z) := 1 + 2* x12.
   
   Definition P_id_proper (x12:Z) := 2* x12.
   
   Definition P_id_U22 (x12:Z) := 3 + 1* x12.
   
   Definition P_id_e  := 2.
   
   Definition P_id_U61 (x12:Z) := 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 1* x12.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_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 := 
     (* <isQid(ok(X_)),isQid(X_)> *)
    | DP_R_xml_0_scc_2_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
       DP_R_xml_0_scc_2 (algebra.Alg.Term algebra.F.id_isQid (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_isQid (x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_2.
   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 (x12:Z) := 2* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_isList (x12:Z) := 2 + 2* x12.
   
   Definition P_id_i  := 2.
   
   Definition P_id_U11 (x12:Z) := 3 + 1* x12.
   
   Definition P_id_isQid (x12:Z) := 2 + 2* x12.
   
   Definition P_id_isNeList (x12:Z) := 2* x12.
   
   Definition P_id_ok (x12:Z) := 1 + 1* x12.
   
   Definition P_id_mark (x12:Z) := 0.
   
   Definition P_id_isPal (x12:Z) := 1* x12.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_u  := 1.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_a  := 1.
   
   Definition P_id_U52 (x12:Z) := 2* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_U72 (x12:Z) := 1* x12.
   
   Definition P_id_U31 (x12:Z) := 2 + 2* x12.
   
   Definition P_id_o  := 1.
   
   Definition P_id_tt  := 2.
   
   Definition P_id_isNePal (x12:Z) := 2* x12.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 2.
   
   Definition P_id_U81 (x12:Z) := 2* x12.
   
   Definition P_id_U42 (x12:Z) := 1 + 2* x12.
   
   Definition P_id_proper (x12:Z) := 2* x12.
   
   Definition P_id_U22 (x12:Z) := 3 + 1* x12.
   
   Definition P_id_e  := 2.
   
   Definition P_id_U61 (x12:Z) := 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 1* x12.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_2.
   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.
  
  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 := 
     (* <isPal(ok(X_)),isPal(X_)> *)
    | DP_R_xml_0_scc_3_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
       DP_R_xml_0_scc_3 (algebra.Alg.Term algebra.F.id_isPal (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_isPal (x12::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 (x12:Z) := 2* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_isList (x12:Z) := 2 + 2* x12.
   
   Definition P_id_i  := 2.
   
   Definition P_id_U11 (x12:Z) := 3 + 1* x12.
   
   Definition P_id_isQid (x12:Z) := 2 + 2* x12.
   
   Definition P_id_isNeList (x12:Z) := 2* x12.
   
   Definition P_id_ok (x12:Z) := 1 + 1* x12.
   
   Definition P_id_mark (x12:Z) := 0.
   
   Definition P_id_isPal (x12:Z) := 1* x12.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_u  := 1.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_a  := 1.
   
   Definition P_id_U52 (x12:Z) := 2* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_U72 (x12:Z) := 1* x12.
   
   Definition P_id_U31 (x12:Z) := 2 + 2* x12.
   
   Definition P_id_o  := 1.
   
   Definition P_id_tt  := 2.
   
   Definition P_id_isNePal (x12:Z) := 2* x12.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 2.
   
   Definition P_id_U81 (x12:Z) := 2* x12.
   
   Definition P_id_U42 (x12:Z) := 1 + 2* x12.
   
   Definition P_id_proper (x12:Z) := 2* x12.
   
   Definition P_id_U22 (x12:Z) := 3 + 1* x12.
   
   Definition P_id_e  := 2.
   
   Definition P_id_U61 (x12:Z) := 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 1* x12.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_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 := 
     (* <isNeList(ok(X_)),isNeList(X_)> *)
    | DP_R_xml_0_scc_4_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
       DP_R_xml_0_scc_4 (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_isNeList (x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_4.
   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 (x12:Z) := 2* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_isList (x12:Z) := 2 + 2* x12.
   
   Definition P_id_i  := 2.
   
   Definition P_id_U11 (x12:Z) := 3 + 1* x12.
   
   Definition P_id_isQid (x12:Z) := 2 + 2* x12.
   
   Definition P_id_isNeList (x12:Z) := 2* x12.
   
   Definition P_id_ok (x12:Z) := 1 + 1* x12.
   
   Definition P_id_mark (x12:Z) := 0.
   
   Definition P_id_isPal (x12:Z) := 1* x12.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_u  := 1.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_a  := 1.
   
   Definition P_id_U52 (x12:Z) := 2* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_U72 (x12:Z) := 1* x12.
   
   Definition P_id_U31 (x12:Z) := 2 + 2* x12.
   
   Definition P_id_o  := 1.
   
   Definition P_id_tt  := 2.
   
   Definition P_id_isNePal (x12:Z) := 2* x12.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 2.
   
   Definition P_id_U81 (x12:Z) := 2* x12.
   
   Definition P_id_U42 (x12:Z) := 1 + 2* x12.
   
   Definition P_id_proper (x12:Z) := 2* x12.
   
   Definition P_id_U22 (x12:Z) := 3 + 1* x12.
   
   Definition P_id_e  := 2.
   
   Definition P_id_U61 (x12:Z) := 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 1* x12.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_4.
   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.
  
  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 := 
     (* <isList(ok(X_)),isList(X_)> *)
    | DP_R_xml_0_scc_5_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
       DP_R_xml_0_scc_5 (algebra.Alg.Term algebra.F.id_isList (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_isList (x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_5.
   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 (x12:Z) := 2* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_isList (x12:Z) := 2 + 2* x12.
   
   Definition P_id_i  := 2.
   
   Definition P_id_U11 (x12:Z) := 3 + 1* x12.
   
   Definition P_id_isQid (x12:Z) := 2 + 2* x12.
   
   Definition P_id_isNeList (x12:Z) := 2* x12.
   
   Definition P_id_ok (x12:Z) := 1 + 1* x12.
   
   Definition P_id_mark (x12:Z) := 0.
   
   Definition P_id_isPal (x12:Z) := 1* x12.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_u  := 1.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_a  := 1.
   
   Definition P_id_U52 (x12:Z) := 2* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_U72 (x12:Z) := 1* x12.
   
   Definition P_id_U31 (x12:Z) := 2 + 2* x12.
   
   Definition P_id_o  := 1.
   
   Definition P_id_tt  := 2.
   
   Definition P_id_isNePal (x12:Z) := 2* x12.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 2.
   
   Definition P_id_U81 (x12:Z) := 2* x12.
   
   Definition P_id_U42 (x12:Z) := 1 + 2* x12.
   
   Definition P_id_proper (x12:Z) := 2* x12.
   
   Definition P_id_U22 (x12:Z) := 3 + 1* x12.
   
   Definition P_id_e  := 2.
   
   Definition P_id_U61 (x12:Z) := 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 1* x12.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_5.
   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.
  
  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_non_scc_6  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(isNePal(X_)),isNePal(proper(X_))> *)
    | DP_R_xml_0_non_scc_6_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isNePal (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_6 (algebra.Alg.Term algebra.F.id_isNePal 
                             ((algebra.Alg.Term algebra.F.id_proper 
                             (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_6 :
   forall x y, 
    (DP_R_xml_0_non_scc_6 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_7  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(isQid(X_)),isQid(proper(X_))> *)
    | DP_R_xml_0_non_scc_7_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isQid (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_7 (algebra.Alg.Term algebra.F.id_isQid 
                             ((algebra.Alg.Term algebra.F.id_proper 
                             (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_7 :
   forall x y, 
    (DP_R_xml_0_non_scc_7 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_8  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(isPal(X_)),isPal(proper(X_))> *)
    | DP_R_xml_0_non_scc_8_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isPal (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_8 (algebra.Alg.Term algebra.F.id_isPal 
                             ((algebra.Alg.Term algebra.F.id_proper 
                             (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_8 :
   forall x y, 
    (DP_R_xml_0_non_scc_8 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_9  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(isNeList(X_)),isNeList(proper(X_))> *)
    | DP_R_xml_0_non_scc_9_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_9 (algebra.Alg.Term algebra.F.id_isNeList 
                             ((algebra.Alg.Term algebra.F.id_proper 
                             (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::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_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_10  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(isList(X_)),isList(proper(X_))> *)
    | DP_R_xml_0_non_scc_10_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_10 (algebra.Alg.Term algebra.F.id_isList 
                              ((algebra.Alg.Term algebra.F.id_proper 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::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_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_scc_11  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <U81(ok(X_)),U81(X_)> *)
    | DP_R_xml_0_scc_11_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
       DP_R_xml_0_scc_11 (algebra.Alg.Term algebra.F.id_U81 (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_U81 (x12::nil))
     (* <U81(mark(X_)),U81(X_)> *)
    | DP_R_xml_0_scc_11_1 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
       DP_R_xml_0_scc_11 (algebra.Alg.Term algebra.F.id_U81 (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_U81 (x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_11.
   Inductive DP_R_xml_0_scc_11_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U81(ok(X_)),U81(X_)> *)
     | DP_R_xml_0_scc_11_large_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
        DP_R_xml_0_scc_11_large (algebra.Alg.Term algebra.F.id_U81 (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_U81 (x12::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_11_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U81(mark(X_)),U81(X_)> *)
     | DP_R_xml_0_scc_11_strict_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
        DP_R_xml_0_scc_11_strict (algebra.Alg.Term algebra.F.id_U81 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_U81 (x12::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_11_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 (x12:Z) := 2* x12.
    
    Definition P_id_U71 (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_isList (x12:Z) := 2 + 2* x12.
    
    Definition P_id_i  := 2.
    
    Definition P_id_U11 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_isQid (x12:Z) := 2 + 2* x12.
    
    Definition P_id_isNeList (x12:Z) := 2* x12.
    
    Definition P_id_ok (x12:Z) := 1 + 1* x12.
    
    Definition P_id_mark (x12:Z) := 0.
    
    Definition P_id_isPal (x12:Z) := 1* x12.
    
    Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
    
    Definition P_id_u  := 1.
    
    Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_a  := 1.
    
    Definition P_id_U52 (x12:Z) := 2* x12.
    
    Definition P_id___ (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_U72 (x12:Z) := 1* x12.
    
    Definition P_id_U31 (x12:Z) := 2 + 2* x12.
    
    Definition P_id_o  := 1.
    
    Definition P_id_tt  := 2.
    
    Definition P_id_isNePal (x12:Z) := 2* x12.
    
    Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
    
    Definition P_id_top (x12:Z) := 0.
    
    Definition P_id_nil  := 2.
    
    Definition P_id_U81 (x12:Z) := 2* x12.
    
    Definition P_id_U42 (x12:Z) := 1 + 2* x12.
    
    Definition P_id_proper (x12:Z) := 2* x12.
    
    Definition P_id_U22 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_e  := 2.
    
    Definition P_id_U61 (x12:Z) := 1* x12.
    
    Lemma P_id_active_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_isList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
    Proof.
      intros x13 x12.
      intros [H_1 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 x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U72_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_top_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
    Proof.
      intros x13 x12.
      intros [H_1 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_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_active x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_i_bounded : 0 <= P_id_i .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
    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 x12, (0 <= x12) ->0 <= P_id_mark x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_u_bounded : 0 <= P_id_u .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_a_bounded : 0 <= P_id_a .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_o_bounded : 0 <= P_id_o .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_tt_bounded : 0 <= P_id_tt .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
    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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_proper x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_e_bounded : 0 <= P_id_e .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
    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_U71 P_id_isList P_id_i P_id_U11 
       P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
       P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
       P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
       P_id_proper P_id_U22 P_id_e P_id_U61.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                     P_id_active (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                     P_id_U71 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                     P_id_isList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                    | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                     P_id_U11 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                     P_id_isQid (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                     P_id_isNeList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                     P_id_ok (measure x12)
                    | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                     P_id_mark (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                     P_id_isPal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                     P_id_U41 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                    | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                     P_id_U21 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                    | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                     P_id_U52 (measure x12)
                    | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                     P_id___ (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                     P_id_U72 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                     P_id_U31 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                    | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                    | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                     P_id_isNePal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                     P_id_U51 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                     P_id_top (measure x12)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                     P_id_U81 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                     P_id_U42 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                     P_id_proper (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                     P_id_U22 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                    | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                     P_id_U61 (measure x12)
                    | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_U22_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNEPAL (x12:Z) := 0.
    
    Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U52_hat_1 (x12:Z) := 0.
    
    Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U42_hat_1 (x12:Z) := 0.
    
    Definition P_id_TOP (x12:Z) := 0.
    
    Definition P_id_ISQID (x12:Z) := 0.
    
    Definition P_id_ISPAL (x12:Z) := 0.
    
    Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_ACTIVE (x12:Z) := 0.
    
    Definition P_id_ISLIST (x12:Z) := 0.
    
    Definition P_id_PROPER (x12:Z) := 0.
    
    Definition P_id_U31_hat_1 (x12:Z) := 0.
    
    Definition P_id_U72_hat_1 (x12:Z) := 0.
    
    Definition P_id_U61_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNELIST (x12:Z) := 0.
    
    Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U11_hat_1 (x12:Z) := 0.
    
    Definition P_id_U81_hat_1 (x12:Z) := 1* x12.
    
    Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
    
    Lemma P_id_U22_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNEPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_TOP_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISQID_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISLIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_PROPER_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNELIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U11_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Definition marked_measure  := 
      InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
       P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
       P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
       P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
       P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
       P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
       P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
       P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
       P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
       P_id_U81_hat_1 P_id____hat_1.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                            P_id_U22_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNePal 
                              (x12::nil)) =>
                            P_id_ISNEPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U21 (x13::
                              x12::nil)) =>
                            P_id_U21_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                            P_id_U52_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U51 (x13::
                              x12::nil)) =>
                            P_id_U51_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                            P_id_U42_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                            P_id_TOP (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                            P_id_ISQID (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                            P_id_ISPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U71 (x13::
                              x12::nil)) =>
                            P_id_U71_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_active 
                              (x12::nil)) =>
                            P_id_ACTIVE (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isList 
                              (x12::nil)) =>
                            P_id_ISLIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_proper 
                              (x12::nil)) =>
                            P_id_PROPER (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                            P_id_U31_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                            P_id_U72_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                            P_id_U61_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNeList 
                              (x12::nil)) =>
                            P_id_ISNELIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U41 (x13::
                              x12::nil)) =>
                            P_id_U41_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                            P_id_U11_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                            P_id_U81_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id___ (x13::
                              x12::nil)) =>
                            P_id____hat_1 (measure x13) (measure x12)
                           | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_U22_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNEPAL_monotonic;assumption.
      intros ;apply P_id_U21_hat_1_monotonic;assumption.
      intros ;apply P_id_U52_hat_1_monotonic;assumption.
      intros ;apply P_id_U51_hat_1_monotonic;assumption.
      intros ;apply P_id_U42_hat_1_monotonic;assumption.
      intros ;apply P_id_TOP_monotonic;assumption.
      intros ;apply P_id_ISQID_monotonic;assumption.
      intros ;apply P_id_ISPAL_monotonic;assumption.
      intros ;apply P_id_U71_hat_1_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_ISLIST_monotonic;assumption.
      intros ;apply P_id_PROPER_monotonic;assumption.
      intros ;apply P_id_U31_hat_1_monotonic;assumption.
      intros ;apply P_id_U72_hat_1_monotonic;assumption.
      intros ;apply P_id_U61_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNELIST_monotonic;assumption.
      intros ;apply P_id_U41_hat_1_monotonic;assumption.
      intros ;apply P_id_U11_hat_1_monotonic;assumption.
      intros ;apply P_id_U81_hat_1_monotonic;assumption.
      intros ;apply P_id____hat_1_monotonic;assumption.
    Qed.
    
    Ltac rewrite_and_unfold  :=
     do 2 (rewrite marked_measure_equation);
      repeat (
      match goal with
        |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
         rewrite (measure_equation (algebra.Alg.Term f t))
        | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
         rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
        end
      ).
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_11.DP_R_xml_0_scc_11_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_11_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 (x12:Z) := 1 + 2* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 2* x12.
   
   Definition P_id_isList (x12:Z) := 0.
   
   Definition P_id_i  := 0.
   
   Definition P_id_U11 (x12:Z) := 1* x12.
   
   Definition P_id_isQid (x12:Z) := 0.
   
   Definition P_id_isNeList (x12:Z) := 0.
   
   Definition P_id_ok (x12:Z) := 1* x12.
   
   Definition P_id_mark (x12:Z) := 1 + 1* x12.
   
   Definition P_id_isPal (x12:Z) := 1.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_u  := 0.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_a  := 0.
   
   Definition P_id_U52 (x12:Z) := 1* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 2 + 1* x12 + 2* x13.
   
   Definition P_id_U72 (x12:Z) := 1 + 1* x12.
   
   Definition P_id_U31 (x12:Z) := 1* x12.
   
   Definition P_id_o  := 2.
   
   Definition P_id_tt  := 0.
   
   Definition P_id_isNePal (x12:Z) := 1.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_U81 (x12:Z) := 1* x12.
   
   Definition P_id_U42 (x12:Z) := 1* x12.
   
   Definition P_id_proper (x12:Z) := 2* x12.
   
   Definition P_id_U22 (x12:Z) := 1* x12.
   
   Definition P_id_e  := 0.
   
   Definition P_id_U61 (x12:Z) := 2 + 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 2* x12.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
   
   Definition le a b := marked_measure a <= marked_measure b.
   
   Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
   Proof.
     unfold lt, le in *.
     intros a b c.
     apply (interp.le_lt_compat_right (interp.o_Z 0)).
   Qed.
   
   Lemma wf_lt : well_founded lt.
   Proof.
     unfold lt in *.
     apply Inverse_Image.wf_inverse_image with  (B:=Z).
     apply Zwf.Zwf_well_founded.
   Qed.
   
   Lemma DP_R_xml_0_scc_11_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_11_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_11_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_11_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_11_large  := WF_DP_R_xml_0_scc_11_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_11.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_11_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_11_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_11_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_11_large.
   Qed.
  End WF_DP_R_xml_0_scc_11.
  
  Definition wf_DP_R_xml_0_scc_11  := WF_DP_R_xml_0_scc_11.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_11 :
   forall x y, (DP_R_xml_0_scc_11 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_11).
    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_11.
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_12  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(U81(X_)),U81(proper(X_))> *)
    | DP_R_xml_0_non_scc_12_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U81 (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_12 (algebra.Alg.Term algebra.F.id_U81 
                              ((algebra.Alg.Term algebra.F.id_proper 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::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_11;
        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_13  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <U72(ok(X_)),U72(X_)> *)
    | DP_R_xml_0_scc_13_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
       DP_R_xml_0_scc_13 (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_U72 (x12::nil))
     (* <U72(mark(X_)),U72(X_)> *)
    | DP_R_xml_0_scc_13_1 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
       DP_R_xml_0_scc_13 (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_U72 (x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_13.
   Inductive DP_R_xml_0_scc_13_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U72(ok(X_)),U72(X_)> *)
     | DP_R_xml_0_scc_13_large_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
        DP_R_xml_0_scc_13_large (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_U72 (x12::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_13_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U72(mark(X_)),U72(X_)> *)
     | DP_R_xml_0_scc_13_strict_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
        DP_R_xml_0_scc_13_strict (algebra.Alg.Term algebra.F.id_U72 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_U72 (x12::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_13_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 (x12:Z) := 2* x12.
    
    Definition P_id_U71 (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_isList (x12:Z) := 2 + 2* x12.
    
    Definition P_id_i  := 2.
    
    Definition P_id_U11 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_isQid (x12:Z) := 2 + 2* x12.
    
    Definition P_id_isNeList (x12:Z) := 2* x12.
    
    Definition P_id_ok (x12:Z) := 1 + 1* x12.
    
    Definition P_id_mark (x12:Z) := 0.
    
    Definition P_id_isPal (x12:Z) := 1* x12.
    
    Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
    
    Definition P_id_u  := 1.
    
    Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_a  := 1.
    
    Definition P_id_U52 (x12:Z) := 2* x12.
    
    Definition P_id___ (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_U72 (x12:Z) := 1* x12.
    
    Definition P_id_U31 (x12:Z) := 2 + 2* x12.
    
    Definition P_id_o  := 1.
    
    Definition P_id_tt  := 2.
    
    Definition P_id_isNePal (x12:Z) := 2* x12.
    
    Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
    
    Definition P_id_top (x12:Z) := 0.
    
    Definition P_id_nil  := 2.
    
    Definition P_id_U81 (x12:Z) := 2* x12.
    
    Definition P_id_U42 (x12:Z) := 1 + 2* x12.
    
    Definition P_id_proper (x12:Z) := 2* x12.
    
    Definition P_id_U22 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_e  := 2.
    
    Definition P_id_U61 (x12:Z) := 1* x12.
    
    Lemma P_id_active_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_isList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
    Proof.
      intros x13 x12.
      intros [H_1 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 x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U72_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_top_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
    Proof.
      intros x13 x12.
      intros [H_1 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_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_active x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_i_bounded : 0 <= P_id_i .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
    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 x12, (0 <= x12) ->0 <= P_id_mark x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_u_bounded : 0 <= P_id_u .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_a_bounded : 0 <= P_id_a .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_o_bounded : 0 <= P_id_o .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_tt_bounded : 0 <= P_id_tt .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
    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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_proper x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_e_bounded : 0 <= P_id_e .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
    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_U71 P_id_isList P_id_i P_id_U11 
       P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
       P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
       P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
       P_id_proper P_id_U22 P_id_e P_id_U61.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                     P_id_active (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                     P_id_U71 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                     P_id_isList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                    | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                     P_id_U11 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                     P_id_isQid (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                     P_id_isNeList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                     P_id_ok (measure x12)
                    | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                     P_id_mark (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                     P_id_isPal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                     P_id_U41 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                    | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                     P_id_U21 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                    | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                     P_id_U52 (measure x12)
                    | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                     P_id___ (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                     P_id_U72 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                     P_id_U31 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                    | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                    | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                     P_id_isNePal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                     P_id_U51 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                     P_id_top (measure x12)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                     P_id_U81 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                     P_id_U42 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                     P_id_proper (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                     P_id_U22 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                    | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                     P_id_U61 (measure x12)
                    | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_U22_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNEPAL (x12:Z) := 0.
    
    Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U52_hat_1 (x12:Z) := 0.
    
    Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U42_hat_1 (x12:Z) := 0.
    
    Definition P_id_TOP (x12:Z) := 0.
    
    Definition P_id_ISQID (x12:Z) := 0.
    
    Definition P_id_ISPAL (x12:Z) := 0.
    
    Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_ACTIVE (x12:Z) := 0.
    
    Definition P_id_ISLIST (x12:Z) := 0.
    
    Definition P_id_PROPER (x12:Z) := 0.
    
    Definition P_id_U31_hat_1 (x12:Z) := 0.
    
    Definition P_id_U72_hat_1 (x12:Z) := 1* x12.
    
    Definition P_id_U61_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNELIST (x12:Z) := 0.
    
    Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U11_hat_1 (x12:Z) := 0.
    
    Definition P_id_U81_hat_1 (x12:Z) := 0.
    
    Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
    
    Lemma P_id_U22_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNEPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_TOP_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISQID_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISLIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_PROPER_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNELIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U11_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Definition marked_measure  := 
      InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
       P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
       P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
       P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
       P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
       P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
       P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
       P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
       P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
       P_id_U81_hat_1 P_id____hat_1.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                            P_id_U22_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNePal 
                              (x12::nil)) =>
                            P_id_ISNEPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U21 (x13::
                              x12::nil)) =>
                            P_id_U21_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                            P_id_U52_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U51 (x13::
                              x12::nil)) =>
                            P_id_U51_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                            P_id_U42_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                            P_id_TOP (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                            P_id_ISQID (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                            P_id_ISPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U71 (x13::
                              x12::nil)) =>
                            P_id_U71_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_active 
                              (x12::nil)) =>
                            P_id_ACTIVE (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isList 
                              (x12::nil)) =>
                            P_id_ISLIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_proper 
                              (x12::nil)) =>
                            P_id_PROPER (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                            P_id_U31_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                            P_id_U72_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                            P_id_U61_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNeList 
                              (x12::nil)) =>
                            P_id_ISNELIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U41 (x13::
                              x12::nil)) =>
                            P_id_U41_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                            P_id_U11_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                            P_id_U81_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id___ (x13::
                              x12::nil)) =>
                            P_id____hat_1 (measure x13) (measure x12)
                           | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_U22_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNEPAL_monotonic;assumption.
      intros ;apply P_id_U21_hat_1_monotonic;assumption.
      intros ;apply P_id_U52_hat_1_monotonic;assumption.
      intros ;apply P_id_U51_hat_1_monotonic;assumption.
      intros ;apply P_id_U42_hat_1_monotonic;assumption.
      intros ;apply P_id_TOP_monotonic;assumption.
      intros ;apply P_id_ISQID_monotonic;assumption.
      intros ;apply P_id_ISPAL_monotonic;assumption.
      intros ;apply P_id_U71_hat_1_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_ISLIST_monotonic;assumption.
      intros ;apply P_id_PROPER_monotonic;assumption.
      intros ;apply P_id_U31_hat_1_monotonic;assumption.
      intros ;apply P_id_U72_hat_1_monotonic;assumption.
      intros ;apply P_id_U61_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNELIST_monotonic;assumption.
      intros ;apply P_id_U41_hat_1_monotonic;assumption.
      intros ;apply P_id_U11_hat_1_monotonic;assumption.
      intros ;apply P_id_U81_hat_1_monotonic;assumption.
      intros ;apply P_id____hat_1_monotonic;assumption.
    Qed.
    
    Ltac rewrite_and_unfold  :=
     do 2 (rewrite marked_measure_equation);
      repeat (
      match goal with
        |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
         rewrite (measure_equation (algebra.Alg.Term f t))
        | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
         rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
        end
      ).
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_13.DP_R_xml_0_scc_13_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_13_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 (x12:Z) := 1 + 2* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 2* x12.
   
   Definition P_id_isList (x12:Z) := 0.
   
   Definition P_id_i  := 0.
   
   Definition P_id_U11 (x12:Z) := 1* x12.
   
   Definition P_id_isQid (x12:Z) := 0.
   
   Definition P_id_isNeList (x12:Z) := 0.
   
   Definition P_id_ok (x12:Z) := 1* x12.
   
   Definition P_id_mark (x12:Z) := 1 + 1* x12.
   
   Definition P_id_isPal (x12:Z) := 1.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_u  := 0.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_a  := 0.
   
   Definition P_id_U52 (x12:Z) := 1* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 2 + 1* x12 + 2* x13.
   
   Definition P_id_U72 (x12:Z) := 1 + 1* x12.
   
   Definition P_id_U31 (x12:Z) := 1* x12.
   
   Definition P_id_o  := 2.
   
   Definition P_id_tt  := 0.
   
   Definition P_id_isNePal (x12:Z) := 1.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_U81 (x12:Z) := 1* x12.
   
   Definition P_id_U42 (x12:Z) := 1* x12.
   
   Definition P_id_proper (x12:Z) := 2* x12.
   
   Definition P_id_U22 (x12:Z) := 1* x12.
   
   Definition P_id_e  := 0.
   
   Definition P_id_U61 (x12:Z) := 2 + 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 2* x12.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
   
   Definition le a b := marked_measure a <= marked_measure b.
   
   Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
   Proof.
     unfold lt, le in *.
     intros a b c.
     apply (interp.le_lt_compat_right (interp.o_Z 0)).
   Qed.
   
   Lemma wf_lt : well_founded lt.
   Proof.
     unfold lt in *.
     apply Inverse_Image.wf_inverse_image with  (B:=Z).
     apply Zwf.Zwf_well_founded.
   Qed.
   
   Lemma DP_R_xml_0_scc_13_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_13_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_13_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_13_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_13_large  := WF_DP_R_xml_0_scc_13_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_13.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_13_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_13_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_13_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_13_large.
   Qed.
  End WF_DP_R_xml_0_scc_13.
  
  Definition wf_DP_R_xml_0_scc_13  := WF_DP_R_xml_0_scc_13.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_13 :
   forall x y, (DP_R_xml_0_scc_13 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_13).
    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_13.
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_14  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(U72(X_)),U72(proper(X_))> *)
    | DP_R_xml_0_non_scc_14_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_14 (algebra.Alg.Term algebra.F.id_U72 
                              ((algebra.Alg.Term algebra.F.id_proper 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::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_13;
        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_15  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <U71(ok(X1_),ok(X2_)),U71(X1_,X2_)> *)
    | DP_R_xml_0_scc_15_0 :
     forall x12 x10 x9 x13, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x9::nil)) x13) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x10::nil)) x12) ->
        DP_R_xml_0_scc_15 (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) 
         (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil))
     (* <U71(mark(X1_),X2_),U71(X1_,X2_)> *)
    | DP_R_xml_0_scc_15_1 :
     forall x12 x10 x9 x13, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x9::nil)) x13) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x10 x12) ->
        DP_R_xml_0_scc_15 (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) 
         (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_15.
   Inductive DP_R_xml_0_scc_15_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U71(mark(X1_),X2_),U71(X1_,X2_)> *)
     | DP_R_xml_0_scc_15_large_0 :
      forall x12 x10 x9 x13, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x9::nil)) x13) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   x10 x12) ->
         DP_R_xml_0_scc_15_large (algebra.Alg.Term algebra.F.id_U71 (x9::
                                  x10::nil)) 
          (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_15_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U71(ok(X1_),ok(X2_)),U71(X1_,X2_)> *)
     | DP_R_xml_0_scc_15_strict_0 :
      forall x12 x10 x9 x13, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x9::nil)) x13) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_ok (x10::nil)) x12) ->
         DP_R_xml_0_scc_15_strict (algebra.Alg.Term algebra.F.id_U71 (x9::
                                   x10::nil)) 
          (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_15_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 (x12:Z) := 2* x12.
    
    Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 3* x12 + 1* x13.
    
    Definition P_id_isList (x12:Z) := 2 + 2* x12.
    
    Definition P_id_i  := 2.
    
    Definition P_id_U11 (x12:Z) := 1 + 1* x12.
    
    Definition P_id_isQid (x12:Z) := 2* x12.
    
    Definition P_id_isNeList (x12:Z) := 2 + 1* x12.
    
    Definition P_id_ok (x12:Z) := 0.
    
    Definition P_id_mark (x12:Z) := 1 + 1* x12.
    
    Definition P_id_isPal (x12:Z) := 3 + 2* x12.
    
    Definition P_id_U41 (x12:Z) (x13:Z) := 1 + 2* x12 + 2* x13.
    
    Definition P_id_u  := 2.
    
    Definition P_id_U21 (x12:Z) (x13:Z) := 2* x12 + 3* x13.
    
    Definition P_id_a  := 2.
    
    Definition P_id_U52 (x12:Z) := 1 + 3* x12.
    
    Definition P_id___ (x12:Z) (x13:Z) := 1 + 3* x12 + 2* x13.
    
    Definition P_id_U72 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_U31 (x12:Z) := 1* x12.
    
    Definition P_id_o  := 2.
    
    Definition P_id_tt  := 2.
    
    Definition P_id_isNePal (x12:Z) := 2 + 2* x12.
    
    Definition P_id_U51 (x12:Z) (x13:Z) := 2* x12 + 3* x13.
    
    Definition P_id_top (x12:Z) := 0.
    
    Definition P_id_nil  := 0.
    
    Definition P_id_U81 (x12:Z) := 2* x12.
    
    Definition P_id_U42 (x12:Z) := 1 + 3* x12.
    
    Definition P_id_proper (x12:Z) := 1* x12.
    
    Definition P_id_U22 (x12:Z) := 1 + 2* x12.
    
    Definition P_id_e  := 2.
    
    Definition P_id_U61 (x12:Z) := 2* x12.
    
    Lemma P_id_active_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_isList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
    Proof.
      intros x13 x12.
      intros [H_1 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 x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U72_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_top_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
    Proof.
      intros x13 x12.
      intros [H_1 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_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_active x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_i_bounded : 0 <= P_id_i .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
    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 x12, (0 <= x12) ->0 <= P_id_mark x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_u_bounded : 0 <= P_id_u .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_a_bounded : 0 <= P_id_a .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_o_bounded : 0 <= P_id_o .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_tt_bounded : 0 <= P_id_tt .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
    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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_proper x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_e_bounded : 0 <= P_id_e .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
    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_U71 P_id_isList P_id_i P_id_U11 
       P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
       P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
       P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
       P_id_proper P_id_U22 P_id_e P_id_U61.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                     P_id_active (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                     P_id_U71 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                     P_id_isList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                    | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                     P_id_U11 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                     P_id_isQid (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                     P_id_isNeList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                     P_id_ok (measure x12)
                    | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                     P_id_mark (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                     P_id_isPal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                     P_id_U41 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                    | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                     P_id_U21 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                    | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                     P_id_U52 (measure x12)
                    | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                     P_id___ (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                     P_id_U72 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                     P_id_U31 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                    | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                    | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                     P_id_isNePal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                     P_id_U51 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                     P_id_top (measure x12)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                     P_id_U81 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                     P_id_U42 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                     P_id_proper (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                     P_id_U22 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                    | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                     P_id_U61 (measure x12)
                    | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_U22_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNEPAL (x12:Z) := 0.
    
    Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U52_hat_1 (x12:Z) := 0.
    
    Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U42_hat_1 (x12:Z) := 0.
    
    Definition P_id_TOP (x12:Z) := 0.
    
    Definition P_id_ISQID (x12:Z) := 0.
    
    Definition P_id_ISPAL (x12:Z) := 0.
    
    Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 1* x12.
    
    Definition P_id_ACTIVE (x12:Z) := 0.
    
    Definition P_id_ISLIST (x12:Z) := 0.
    
    Definition P_id_PROPER (x12:Z) := 0.
    
    Definition P_id_U31_hat_1 (x12:Z) := 0.
    
    Definition P_id_U72_hat_1 (x12:Z) := 0.
    
    Definition P_id_U61_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNELIST (x12:Z) := 0.
    
    Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U11_hat_1 (x12:Z) := 0.
    
    Definition P_id_U81_hat_1 (x12:Z) := 0.
    
    Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
    
    Lemma P_id_U22_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNEPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_TOP_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISQID_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISLIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_PROPER_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNELIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U11_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Definition marked_measure  := 
      InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
       P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
       P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
       P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
       P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
       P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
       P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
       P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
       P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
       P_id_U81_hat_1 P_id____hat_1.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                            P_id_U22_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNePal 
                              (x12::nil)) =>
                            P_id_ISNEPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U21 (x13::
                              x12::nil)) =>
                            P_id_U21_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                            P_id_U52_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U51 (x13::
                              x12::nil)) =>
                            P_id_U51_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                            P_id_U42_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                            P_id_TOP (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                            P_id_ISQID (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                            P_id_ISPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U71 (x13::
                              x12::nil)) =>
                            P_id_U71_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_active 
                              (x12::nil)) =>
                            P_id_ACTIVE (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isList 
                              (x12::nil)) =>
                            P_id_ISLIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_proper 
                              (x12::nil)) =>
                            P_id_PROPER (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                            P_id_U31_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                            P_id_U72_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                            P_id_U61_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNeList 
                              (x12::nil)) =>
                            P_id_ISNELIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U41 (x13::
                              x12::nil)) =>
                            P_id_U41_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                            P_id_U11_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                            P_id_U81_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id___ (x13::
                              x12::nil)) =>
                            P_id____hat_1 (measure x13) (measure x12)
                           | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_U22_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNEPAL_monotonic;assumption.
      intros ;apply P_id_U21_hat_1_monotonic;assumption.
      intros ;apply P_id_U52_hat_1_monotonic;assumption.
      intros ;apply P_id_U51_hat_1_monotonic;assumption.
      intros ;apply P_id_U42_hat_1_monotonic;assumption.
      intros ;apply P_id_TOP_monotonic;assumption.
      intros ;apply P_id_ISQID_monotonic;assumption.
      intros ;apply P_id_ISPAL_monotonic;assumption.
      intros ;apply P_id_U71_hat_1_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_ISLIST_monotonic;assumption.
      intros ;apply P_id_PROPER_monotonic;assumption.
      intros ;apply P_id_U31_hat_1_monotonic;assumption.
      intros ;apply P_id_U72_hat_1_monotonic;assumption.
      intros ;apply P_id_U61_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNELIST_monotonic;assumption.
      intros ;apply P_id_U41_hat_1_monotonic;assumption.
      intros ;apply P_id_U11_hat_1_monotonic;assumption.
      intros ;apply P_id_U81_hat_1_monotonic;assumption.
      intros ;apply P_id____hat_1_monotonic;assumption.
    Qed.
    
    Ltac rewrite_and_unfold  :=
     do 2 (rewrite marked_measure_equation);
      repeat (
      match goal with
        |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
         rewrite (measure_equation (algebra.Alg.Term f t))
        | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
         rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
        end
      ).
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_15.DP_R_xml_0_scc_15_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_15_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 (x12:Z) := 1* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 3* x12 + 3* x13.
   
   Definition P_id_isList (x12:Z) := 2* x12.
   
   Definition P_id_i  := 1.
   
   Definition P_id_U11 (x12:Z) := 1* x12.
   
   Definition P_id_isQid (x12:Z) := 2* x12.
   
   Definition P_id_isNeList (x12:Z) := 1* x12.
   
   Definition P_id_ok (x12:Z) := 1 + 2* x12.
   
   Definition P_id_mark (x12:Z) := 0.
   
   Definition P_id_isPal (x12:Z) := 1* x12.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_u  := 3.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_a  := 1.
   
   Definition P_id_U52 (x12:Z) := 3* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_U72 (x12:Z) := 2* x12.
   
   Definition P_id_U31 (x12:Z) := 2* x12.
   
   Definition P_id_o  := 1.
   
   Definition P_id_tt  := 3.
   
   Definition P_id_isNePal (x12:Z) := 2* x12.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 2.
   
   Definition P_id_U81 (x12:Z) := 1* x12.
   
   Definition P_id_U42 (x12:Z) := 2* x12.
   
   Definition P_id_proper (x12:Z) := 3* x12.
   
   Definition P_id_U22 (x12:Z) := 1* x12.
   
   Definition P_id_e  := 1.
   
   Definition P_id_U61 (x12:Z) := 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
   
   Definition le a b := marked_measure a <= marked_measure b.
   
   Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
   Proof.
     unfold lt, le in *.
     intros a b c.
     apply (interp.le_lt_compat_right (interp.o_Z 0)).
   Qed.
   
   Lemma wf_lt : well_founded lt.
   Proof.
     unfold lt in *.
     apply Inverse_Image.wf_inverse_image with  (B:=Z).
     apply Zwf.Zwf_well_founded.
   Qed.
   
   Lemma DP_R_xml_0_scc_15_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_15_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_15_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_15_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_15_large  := WF_DP_R_xml_0_scc_15_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_15.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_15_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_15_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_15_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_15_large.
   Qed.
  End WF_DP_R_xml_0_scc_15.
  
  Definition wf_DP_R_xml_0_scc_15  := WF_DP_R_xml_0_scc_15.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_15 :
   forall x y, (DP_R_xml_0_scc_15 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_15).
    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_15.
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_16  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(U71(X1_,X2_)),U71(proper(X1_),proper(X2_))> *)
    | DP_R_xml_0_non_scc_16_0 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x12) ->
       DP_R_xml_0_non_scc_16 (algebra.Alg.Term algebra.F.id_U71 
                              ((algebra.Alg.Term algebra.F.id_proper 
                              (x9::nil))::(algebra.Alg.Term 
                              algebra.F.id_proper (x10::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::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_15;
        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_17  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <U61(ok(X_)),U61(X_)> *)
    | DP_R_xml_0_scc_17_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
       DP_R_xml_0_scc_17 (algebra.Alg.Term algebra.F.id_U61 (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_U61 (x12::nil))
     (* <U61(mark(X_)),U61(X_)> *)
    | DP_R_xml_0_scc_17_1 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
       DP_R_xml_0_scc_17 (algebra.Alg.Term algebra.F.id_U61 (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_U61 (x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_17.
   Inductive DP_R_xml_0_scc_17_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U61(ok(X_)),U61(X_)> *)
     | DP_R_xml_0_scc_17_large_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
        DP_R_xml_0_scc_17_large (algebra.Alg.Term algebra.F.id_U61 (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_U61 (x12::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_17_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U61(mark(X_)),U61(X_)> *)
     | DP_R_xml_0_scc_17_strict_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
        DP_R_xml_0_scc_17_strict (algebra.Alg.Term algebra.F.id_U61 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_U61 (x12::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_17_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 (x12:Z) := 2* x12.
    
    Definition P_id_U71 (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_isList (x12:Z) := 2 + 2* x12.
    
    Definition P_id_i  := 2.
    
    Definition P_id_U11 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_isQid (x12:Z) := 2 + 2* x12.
    
    Definition P_id_isNeList (x12:Z) := 2* x12.
    
    Definition P_id_ok (x12:Z) := 1 + 1* x12.
    
    Definition P_id_mark (x12:Z) := 0.
    
    Definition P_id_isPal (x12:Z) := 1* x12.
    
    Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
    
    Definition P_id_u  := 1.
    
    Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_a  := 1.
    
    Definition P_id_U52 (x12:Z) := 2* x12.
    
    Definition P_id___ (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_U72 (x12:Z) := 1* x12.
    
    Definition P_id_U31 (x12:Z) := 2 + 2* x12.
    
    Definition P_id_o  := 1.
    
    Definition P_id_tt  := 2.
    
    Definition P_id_isNePal (x12:Z) := 2* x12.
    
    Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
    
    Definition P_id_top (x12:Z) := 0.
    
    Definition P_id_nil  := 2.
    
    Definition P_id_U81 (x12:Z) := 2* x12.
    
    Definition P_id_U42 (x12:Z) := 1 + 2* x12.
    
    Definition P_id_proper (x12:Z) := 2* x12.
    
    Definition P_id_U22 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_e  := 2.
    
    Definition P_id_U61 (x12:Z) := 1* x12.
    
    Lemma P_id_active_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_isList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
    Proof.
      intros x13 x12.
      intros [H_1 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 x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U72_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_top_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
    Proof.
      intros x13 x12.
      intros [H_1 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_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_active x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_i_bounded : 0 <= P_id_i .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
    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 x12, (0 <= x12) ->0 <= P_id_mark x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_u_bounded : 0 <= P_id_u .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_a_bounded : 0 <= P_id_a .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_o_bounded : 0 <= P_id_o .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_tt_bounded : 0 <= P_id_tt .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
    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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_proper x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_e_bounded : 0 <= P_id_e .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
    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_U71 P_id_isList P_id_i P_id_U11 
       P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
       P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
       P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
       P_id_proper P_id_U22 P_id_e P_id_U61.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                     P_id_active (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                     P_id_U71 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                     P_id_isList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                    | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                     P_id_U11 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                     P_id_isQid (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                     P_id_isNeList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                     P_id_ok (measure x12)
                    | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                     P_id_mark (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                     P_id_isPal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                     P_id_U41 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                    | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                     P_id_U21 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                    | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                     P_id_U52 (measure x12)
                    | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                     P_id___ (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                     P_id_U72 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                     P_id_U31 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                    | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                    | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                     P_id_isNePal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                     P_id_U51 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                     P_id_top (measure x12)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                     P_id_U81 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                     P_id_U42 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                     P_id_proper (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                     P_id_U22 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                    | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                     P_id_U61 (measure x12)
                    | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_U22_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNEPAL (x12:Z) := 0.
    
    Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U52_hat_1 (x12:Z) := 0.
    
    Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U42_hat_1 (x12:Z) := 0.
    
    Definition P_id_TOP (x12:Z) := 0.
    
    Definition P_id_ISQID (x12:Z) := 0.
    
    Definition P_id_ISPAL (x12:Z) := 0.
    
    Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_ACTIVE (x12:Z) := 0.
    
    Definition P_id_ISLIST (x12:Z) := 0.
    
    Definition P_id_PROPER (x12:Z) := 0.
    
    Definition P_id_U31_hat_1 (x12:Z) := 0.
    
    Definition P_id_U72_hat_1 (x12:Z) := 0.
    
    Definition P_id_U61_hat_1 (x12:Z) := 1* x12.
    
    Definition P_id_ISNELIST (x12:Z) := 0.
    
    Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U11_hat_1 (x12:Z) := 0.
    
    Definition P_id_U81_hat_1 (x12:Z) := 0.
    
    Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
    
    Lemma P_id_U22_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNEPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_TOP_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISQID_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISLIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_PROPER_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNELIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U11_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Definition marked_measure  := 
      InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
       P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
       P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
       P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
       P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
       P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
       P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
       P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
       P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
       P_id_U81_hat_1 P_id____hat_1.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                            P_id_U22_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNePal 
                              (x12::nil)) =>
                            P_id_ISNEPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U21 (x13::
                              x12::nil)) =>
                            P_id_U21_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                            P_id_U52_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U51 (x13::
                              x12::nil)) =>
                            P_id_U51_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                            P_id_U42_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                            P_id_TOP (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                            P_id_ISQID (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                            P_id_ISPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U71 (x13::
                              x12::nil)) =>
                            P_id_U71_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_active 
                              (x12::nil)) =>
                            P_id_ACTIVE (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isList 
                              (x12::nil)) =>
                            P_id_ISLIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_proper 
                              (x12::nil)) =>
                            P_id_PROPER (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                            P_id_U31_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                            P_id_U72_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                            P_id_U61_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNeList 
                              (x12::nil)) =>
                            P_id_ISNELIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U41 (x13::
                              x12::nil)) =>
                            P_id_U41_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                            P_id_U11_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                            P_id_U81_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id___ (x13::
                              x12::nil)) =>
                            P_id____hat_1 (measure x13) (measure x12)
                           | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_U22_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNEPAL_monotonic;assumption.
      intros ;apply P_id_U21_hat_1_monotonic;assumption.
      intros ;apply P_id_U52_hat_1_monotonic;assumption.
      intros ;apply P_id_U51_hat_1_monotonic;assumption.
      intros ;apply P_id_U42_hat_1_monotonic;assumption.
      intros ;apply P_id_TOP_monotonic;assumption.
      intros ;apply P_id_ISQID_monotonic;assumption.
      intros ;apply P_id_ISPAL_monotonic;assumption.
      intros ;apply P_id_U71_hat_1_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_ISLIST_monotonic;assumption.
      intros ;apply P_id_PROPER_monotonic;assumption.
      intros ;apply P_id_U31_hat_1_monotonic;assumption.
      intros ;apply P_id_U72_hat_1_monotonic;assumption.
      intros ;apply P_id_U61_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNELIST_monotonic;assumption.
      intros ;apply P_id_U41_hat_1_monotonic;assumption.
      intros ;apply P_id_U11_hat_1_monotonic;assumption.
      intros ;apply P_id_U81_hat_1_monotonic;assumption.
      intros ;apply P_id____hat_1_monotonic;assumption.
    Qed.
    
    Ltac rewrite_and_unfold  :=
     do 2 (rewrite marked_measure_equation);
      repeat (
      match goal with
        |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
         rewrite (measure_equation (algebra.Alg.Term f t))
        | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
         rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
        end
      ).
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_17.DP_R_xml_0_scc_17_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_17_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 (x12:Z) := 1 + 2* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 2* x12.
   
   Definition P_id_isList (x12:Z) := 0.
   
   Definition P_id_i  := 0.
   
   Definition P_id_U11 (x12:Z) := 1* x12.
   
   Definition P_id_isQid (x12:Z) := 0.
   
   Definition P_id_isNeList (x12:Z) := 0.
   
   Definition P_id_ok (x12:Z) := 1* x12.
   
   Definition P_id_mark (x12:Z) := 1 + 1* x12.
   
   Definition P_id_isPal (x12:Z) := 1.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_u  := 0.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_a  := 0.
   
   Definition P_id_U52 (x12:Z) := 1* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 2 + 1* x12 + 2* x13.
   
   Definition P_id_U72 (x12:Z) := 1 + 1* x12.
   
   Definition P_id_U31 (x12:Z) := 1* x12.
   
   Definition P_id_o  := 2.
   
   Definition P_id_tt  := 0.
   
   Definition P_id_isNePal (x12:Z) := 1.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_U81 (x12:Z) := 1* x12.
   
   Definition P_id_U42 (x12:Z) := 1* x12.
   
   Definition P_id_proper (x12:Z) := 2* x12.
   
   Definition P_id_U22 (x12:Z) := 1* x12.
   
   Definition P_id_e  := 0.
   
   Definition P_id_U61 (x12:Z) := 2 + 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 2* x12.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
   
   Definition le a b := marked_measure a <= marked_measure b.
   
   Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
   Proof.
     unfold lt, le in *.
     intros a b c.
     apply (interp.le_lt_compat_right (interp.o_Z 0)).
   Qed.
   
   Lemma wf_lt : well_founded lt.
   Proof.
     unfold lt in *.
     apply Inverse_Image.wf_inverse_image with  (B:=Z).
     apply Zwf.Zwf_well_founded.
   Qed.
   
   Lemma DP_R_xml_0_scc_17_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_17_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_17_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_17_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_17_large  := WF_DP_R_xml_0_scc_17_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_17.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_17_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_17_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_17_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_17_large.
   Qed.
  End WF_DP_R_xml_0_scc_17.
  
  Definition wf_DP_R_xml_0_scc_17  := WF_DP_R_xml_0_scc_17.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_17 :
   forall x y, (DP_R_xml_0_scc_17 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_17).
    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_17.
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_18  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(U61(X_)),U61(proper(X_))> *)
    | DP_R_xml_0_non_scc_18_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U61 (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_18 (algebra.Alg.Term algebra.F.id_U61 
                              ((algebra.Alg.Term algebra.F.id_proper 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::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_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_scc_19  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <U52(ok(X_)),U52(X_)> *)
    | DP_R_xml_0_scc_19_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
       DP_R_xml_0_scc_19 (algebra.Alg.Term algebra.F.id_U52 (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_U52 (x12::nil))
     (* <U52(mark(X_)),U52(X_)> *)
    | DP_R_xml_0_scc_19_1 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
       DP_R_xml_0_scc_19 (algebra.Alg.Term algebra.F.id_U52 (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_U52 (x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_19.
   Inductive DP_R_xml_0_scc_19_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U52(ok(X_)),U52(X_)> *)
     | DP_R_xml_0_scc_19_large_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
        DP_R_xml_0_scc_19_large (algebra.Alg.Term algebra.F.id_U52 (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_U52 (x12::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_19_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U52(mark(X_)),U52(X_)> *)
     | DP_R_xml_0_scc_19_strict_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
        DP_R_xml_0_scc_19_strict (algebra.Alg.Term algebra.F.id_U52 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_U52 (x12::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_19_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 (x12:Z) := 2* x12.
    
    Definition P_id_U71 (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_isList (x12:Z) := 2 + 2* x12.
    
    Definition P_id_i  := 2.
    
    Definition P_id_U11 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_isQid (x12:Z) := 2 + 2* x12.
    
    Definition P_id_isNeList (x12:Z) := 2* x12.
    
    Definition P_id_ok (x12:Z) := 1 + 1* x12.
    
    Definition P_id_mark (x12:Z) := 0.
    
    Definition P_id_isPal (x12:Z) := 1* x12.
    
    Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
    
    Definition P_id_u  := 1.
    
    Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_a  := 1.
    
    Definition P_id_U52 (x12:Z) := 2* x12.
    
    Definition P_id___ (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_U72 (x12:Z) := 1* x12.
    
    Definition P_id_U31 (x12:Z) := 2 + 2* x12.
    
    Definition P_id_o  := 1.
    
    Definition P_id_tt  := 2.
    
    Definition P_id_isNePal (x12:Z) := 2* x12.
    
    Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
    
    Definition P_id_top (x12:Z) := 0.
    
    Definition P_id_nil  := 2.
    
    Definition P_id_U81 (x12:Z) := 2* x12.
    
    Definition P_id_U42 (x12:Z) := 1 + 2* x12.
    
    Definition P_id_proper (x12:Z) := 2* x12.
    
    Definition P_id_U22 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_e  := 2.
    
    Definition P_id_U61 (x12:Z) := 1* x12.
    
    Lemma P_id_active_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_isList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
    Proof.
      intros x13 x12.
      intros [H_1 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 x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U72_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_top_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
    Proof.
      intros x13 x12.
      intros [H_1 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_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_active x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_i_bounded : 0 <= P_id_i .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
    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 x12, (0 <= x12) ->0 <= P_id_mark x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_u_bounded : 0 <= P_id_u .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_a_bounded : 0 <= P_id_a .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_o_bounded : 0 <= P_id_o .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_tt_bounded : 0 <= P_id_tt .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
    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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_proper x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_e_bounded : 0 <= P_id_e .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
    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_U71 P_id_isList P_id_i P_id_U11 
       P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
       P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
       P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
       P_id_proper P_id_U22 P_id_e P_id_U61.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                     P_id_active (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                     P_id_U71 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                     P_id_isList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                    | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                     P_id_U11 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                     P_id_isQid (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                     P_id_isNeList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                     P_id_ok (measure x12)
                    | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                     P_id_mark (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                     P_id_isPal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                     P_id_U41 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                    | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                     P_id_U21 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                    | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                     P_id_U52 (measure x12)
                    | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                     P_id___ (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                     P_id_U72 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                     P_id_U31 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                    | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                    | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                     P_id_isNePal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                     P_id_U51 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                     P_id_top (measure x12)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                     P_id_U81 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                     P_id_U42 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                     P_id_proper (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                     P_id_U22 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                    | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                     P_id_U61 (measure x12)
                    | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_U22_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNEPAL (x12:Z) := 0.
    
    Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U52_hat_1 (x12:Z) := 1* x12.
    
    Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U42_hat_1 (x12:Z) := 0.
    
    Definition P_id_TOP (x12:Z) := 0.
    
    Definition P_id_ISQID (x12:Z) := 0.
    
    Definition P_id_ISPAL (x12:Z) := 0.
    
    Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_ACTIVE (x12:Z) := 0.
    
    Definition P_id_ISLIST (x12:Z) := 0.
    
    Definition P_id_PROPER (x12:Z) := 0.
    
    Definition P_id_U31_hat_1 (x12:Z) := 0.
    
    Definition P_id_U72_hat_1 (x12:Z) := 0.
    
    Definition P_id_U61_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNELIST (x12:Z) := 0.
    
    Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U11_hat_1 (x12:Z) := 0.
    
    Definition P_id_U81_hat_1 (x12:Z) := 0.
    
    Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
    
    Lemma P_id_U22_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNEPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_TOP_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISQID_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISLIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_PROPER_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNELIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U11_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Definition marked_measure  := 
      InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
       P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
       P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
       P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
       P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
       P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
       P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
       P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
       P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
       P_id_U81_hat_1 P_id____hat_1.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                            P_id_U22_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNePal 
                              (x12::nil)) =>
                            P_id_ISNEPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U21 (x13::
                              x12::nil)) =>
                            P_id_U21_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                            P_id_U52_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U51 (x13::
                              x12::nil)) =>
                            P_id_U51_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                            P_id_U42_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                            P_id_TOP (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                            P_id_ISQID (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                            P_id_ISPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U71 (x13::
                              x12::nil)) =>
                            P_id_U71_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_active 
                              (x12::nil)) =>
                            P_id_ACTIVE (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isList 
                              (x12::nil)) =>
                            P_id_ISLIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_proper 
                              (x12::nil)) =>
                            P_id_PROPER (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                            P_id_U31_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                            P_id_U72_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                            P_id_U61_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNeList 
                              (x12::nil)) =>
                            P_id_ISNELIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U41 (x13::
                              x12::nil)) =>
                            P_id_U41_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                            P_id_U11_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                            P_id_U81_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id___ (x13::
                              x12::nil)) =>
                            P_id____hat_1 (measure x13) (measure x12)
                           | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_U22_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNEPAL_monotonic;assumption.
      intros ;apply P_id_U21_hat_1_monotonic;assumption.
      intros ;apply P_id_U52_hat_1_monotonic;assumption.
      intros ;apply P_id_U51_hat_1_monotonic;assumption.
      intros ;apply P_id_U42_hat_1_monotonic;assumption.
      intros ;apply P_id_TOP_monotonic;assumption.
      intros ;apply P_id_ISQID_monotonic;assumption.
      intros ;apply P_id_ISPAL_monotonic;assumption.
      intros ;apply P_id_U71_hat_1_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_ISLIST_monotonic;assumption.
      intros ;apply P_id_PROPER_monotonic;assumption.
      intros ;apply P_id_U31_hat_1_monotonic;assumption.
      intros ;apply P_id_U72_hat_1_monotonic;assumption.
      intros ;apply P_id_U61_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNELIST_monotonic;assumption.
      intros ;apply P_id_U41_hat_1_monotonic;assumption.
      intros ;apply P_id_U11_hat_1_monotonic;assumption.
      intros ;apply P_id_U81_hat_1_monotonic;assumption.
      intros ;apply P_id____hat_1_monotonic;assumption.
    Qed.
    
    Ltac rewrite_and_unfold  :=
     do 2 (rewrite marked_measure_equation);
      repeat (
      match goal with
        |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
         rewrite (measure_equation (algebra.Alg.Term f t))
        | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
         rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
        end
      ).
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_19.DP_R_xml_0_scc_19_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_19_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 (x12:Z) := 1 + 2* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 2* x12.
   
   Definition P_id_isList (x12:Z) := 0.
   
   Definition P_id_i  := 0.
   
   Definition P_id_U11 (x12:Z) := 1* x12.
   
   Definition P_id_isQid (x12:Z) := 0.
   
   Definition P_id_isNeList (x12:Z) := 0.
   
   Definition P_id_ok (x12:Z) := 1* x12.
   
   Definition P_id_mark (x12:Z) := 1 + 1* x12.
   
   Definition P_id_isPal (x12:Z) := 1.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_u  := 0.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_a  := 0.
   
   Definition P_id_U52 (x12:Z) := 1* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 2 + 1* x12 + 2* x13.
   
   Definition P_id_U72 (x12:Z) := 1 + 1* x12.
   
   Definition P_id_U31 (x12:Z) := 1* x12.
   
   Definition P_id_o  := 2.
   
   Definition P_id_tt  := 0.
   
   Definition P_id_isNePal (x12:Z) := 1.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_U81 (x12:Z) := 1* x12.
   
   Definition P_id_U42 (x12:Z) := 1* x12.
   
   Definition P_id_proper (x12:Z) := 2* x12.
   
   Definition P_id_U22 (x12:Z) := 1* x12.
   
   Definition P_id_e  := 0.
   
   Definition P_id_U61 (x12:Z) := 2 + 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 2* x12.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
   
   Definition le a b := marked_measure a <= marked_measure b.
   
   Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
   Proof.
     unfold lt, le in *.
     intros a b c.
     apply (interp.le_lt_compat_right (interp.o_Z 0)).
   Qed.
   
   Lemma wf_lt : well_founded lt.
   Proof.
     unfold lt in *.
     apply Inverse_Image.wf_inverse_image with  (B:=Z).
     apply Zwf.Zwf_well_founded.
   Qed.
   
   Lemma DP_R_xml_0_scc_19_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_19_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_19_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_19_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_19_large  := WF_DP_R_xml_0_scc_19_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_19.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_19_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_19_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_19_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_19_large.
   Qed.
  End WF_DP_R_xml_0_scc_19.
  
  Definition wf_DP_R_xml_0_scc_19  := WF_DP_R_xml_0_scc_19.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_19 :
   forall x y, (DP_R_xml_0_scc_19 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_19).
    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_19.
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_20  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(U52(X_)),U52(proper(X_))> *)
    | DP_R_xml_0_non_scc_20_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U52 (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_20 (algebra.Alg.Term algebra.F.id_U52 
                              ((algebra.Alg.Term algebra.F.id_proper 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::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_19;
        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_21  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <U51(ok(X1_),ok(X2_)),U51(X1_,X2_)> *)
    | DP_R_xml_0_scc_21_0 :
     forall x12 x10 x9 x13, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x9::nil)) x13) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x10::nil)) x12) ->
        DP_R_xml_0_scc_21 (algebra.Alg.Term algebra.F.id_U51 (x9::x10::nil)) 
         (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil))
     (* <U51(mark(X1_),X2_),U51(X1_,X2_)> *)
    | DP_R_xml_0_scc_21_1 :
     forall x12 x10 x9 x13, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x9::nil)) x13) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x10 x12) ->
        DP_R_xml_0_scc_21 (algebra.Alg.Term algebra.F.id_U51 (x9::x10::nil)) 
         (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_21.
   Inductive DP_R_xml_0_scc_21_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U51(mark(X1_),X2_),U51(X1_,X2_)> *)
     | DP_R_xml_0_scc_21_large_0 :
      forall x12 x10 x9 x13, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x9::nil)) x13) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   x10 x12) ->
         DP_R_xml_0_scc_21_large (algebra.Alg.Term algebra.F.id_U51 (x9::
                                  x10::nil)) 
          (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_21_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U51(ok(X1_),ok(X2_)),U51(X1_,X2_)> *)
     | DP_R_xml_0_scc_21_strict_0 :
      forall x12 x10 x9 x13, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x9::nil)) x13) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_ok (x10::nil)) x12) ->
         DP_R_xml_0_scc_21_strict (algebra.Alg.Term algebra.F.id_U51 (x9::
                                   x10::nil)) 
          (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_21_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 (x12:Z) := 2* x12.
    
    Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 3* x12 + 1* x13.
    
    Definition P_id_isList (x12:Z) := 2 + 2* x12.
    
    Definition P_id_i  := 2.
    
    Definition P_id_U11 (x12:Z) := 1 + 1* x12.
    
    Definition P_id_isQid (x12:Z) := 2* x12.
    
    Definition P_id_isNeList (x12:Z) := 2 + 1* x12.
    
    Definition P_id_ok (x12:Z) := 0.
    
    Definition P_id_mark (x12:Z) := 1 + 1* x12.
    
    Definition P_id_isPal (x12:Z) := 3 + 2* x12.
    
    Definition P_id_U41 (x12:Z) (x13:Z) := 1 + 2* x12 + 2* x13.
    
    Definition P_id_u  := 2.
    
    Definition P_id_U21 (x12:Z) (x13:Z) := 2* x12 + 3* x13.
    
    Definition P_id_a  := 2.
    
    Definition P_id_U52 (x12:Z) := 1 + 3* x12.
    
    Definition P_id___ (x12:Z) (x13:Z) := 1 + 3* x12 + 2* x13.
    
    Definition P_id_U72 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_U31 (x12:Z) := 1* x12.
    
    Definition P_id_o  := 2.
    
    Definition P_id_tt  := 2.
    
    Definition P_id_isNePal (x12:Z) := 2 + 2* x12.
    
    Definition P_id_U51 (x12:Z) (x13:Z) := 2* x12 + 3* x13.
    
    Definition P_id_top (x12:Z) := 0.
    
    Definition P_id_nil  := 0.
    
    Definition P_id_U81 (x12:Z) := 2* x12.
    
    Definition P_id_U42 (x12:Z) := 1 + 3* x12.
    
    Definition P_id_proper (x12:Z) := 1* x12.
    
    Definition P_id_U22 (x12:Z) := 1 + 2* x12.
    
    Definition P_id_e  := 2.
    
    Definition P_id_U61 (x12:Z) := 2* x12.
    
    Lemma P_id_active_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_isList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
    Proof.
      intros x13 x12.
      intros [H_1 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 x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U72_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_top_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
    Proof.
      intros x13 x12.
      intros [H_1 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_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_active x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_i_bounded : 0 <= P_id_i .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
    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 x12, (0 <= x12) ->0 <= P_id_mark x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_u_bounded : 0 <= P_id_u .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_a_bounded : 0 <= P_id_a .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_o_bounded : 0 <= P_id_o .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_tt_bounded : 0 <= P_id_tt .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
    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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_proper x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_e_bounded : 0 <= P_id_e .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
    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_U71 P_id_isList P_id_i P_id_U11 
       P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
       P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
       P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
       P_id_proper P_id_U22 P_id_e P_id_U61.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                     P_id_active (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                     P_id_U71 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                     P_id_isList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                    | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                     P_id_U11 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                     P_id_isQid (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                     P_id_isNeList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                     P_id_ok (measure x12)
                    | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                     P_id_mark (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                     P_id_isPal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                     P_id_U41 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                    | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                     P_id_U21 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                    | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                     P_id_U52 (measure x12)
                    | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                     P_id___ (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                     P_id_U72 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                     P_id_U31 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                    | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                    | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                     P_id_isNePal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                     P_id_U51 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                     P_id_top (measure x12)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                     P_id_U81 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                     P_id_U42 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                     P_id_proper (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                     P_id_U22 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                    | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                     P_id_U61 (measure x12)
                    | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_U22_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNEPAL (x12:Z) := 0.
    
    Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U52_hat_1 (x12:Z) := 0.
    
    Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 1* x12.
    
    Definition P_id_U42_hat_1 (x12:Z) := 0.
    
    Definition P_id_TOP (x12:Z) := 0.
    
    Definition P_id_ISQID (x12:Z) := 0.
    
    Definition P_id_ISPAL (x12:Z) := 0.
    
    Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_ACTIVE (x12:Z) := 0.
    
    Definition P_id_ISLIST (x12:Z) := 0.
    
    Definition P_id_PROPER (x12:Z) := 0.
    
    Definition P_id_U31_hat_1 (x12:Z) := 0.
    
    Definition P_id_U72_hat_1 (x12:Z) := 0.
    
    Definition P_id_U61_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNELIST (x12:Z) := 0.
    
    Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U11_hat_1 (x12:Z) := 0.
    
    Definition P_id_U81_hat_1 (x12:Z) := 0.
    
    Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
    
    Lemma P_id_U22_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNEPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_TOP_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISQID_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISLIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_PROPER_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNELIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U11_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Definition marked_measure  := 
      InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
       P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
       P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
       P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
       P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
       P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
       P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
       P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
       P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
       P_id_U81_hat_1 P_id____hat_1.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                            P_id_U22_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNePal 
                              (x12::nil)) =>
                            P_id_ISNEPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U21 (x13::
                              x12::nil)) =>
                            P_id_U21_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                            P_id_U52_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U51 (x13::
                              x12::nil)) =>
                            P_id_U51_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                            P_id_U42_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                            P_id_TOP (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                            P_id_ISQID (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                            P_id_ISPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U71 (x13::
                              x12::nil)) =>
                            P_id_U71_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_active 
                              (x12::nil)) =>
                            P_id_ACTIVE (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isList 
                              (x12::nil)) =>
                            P_id_ISLIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_proper 
                              (x12::nil)) =>
                            P_id_PROPER (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                            P_id_U31_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                            P_id_U72_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                            P_id_U61_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNeList 
                              (x12::nil)) =>
                            P_id_ISNELIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U41 (x13::
                              x12::nil)) =>
                            P_id_U41_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                            P_id_U11_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                            P_id_U81_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id___ (x13::
                              x12::nil)) =>
                            P_id____hat_1 (measure x13) (measure x12)
                           | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_U22_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNEPAL_monotonic;assumption.
      intros ;apply P_id_U21_hat_1_monotonic;assumption.
      intros ;apply P_id_U52_hat_1_monotonic;assumption.
      intros ;apply P_id_U51_hat_1_monotonic;assumption.
      intros ;apply P_id_U42_hat_1_monotonic;assumption.
      intros ;apply P_id_TOP_monotonic;assumption.
      intros ;apply P_id_ISQID_monotonic;assumption.
      intros ;apply P_id_ISPAL_monotonic;assumption.
      intros ;apply P_id_U71_hat_1_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_ISLIST_monotonic;assumption.
      intros ;apply P_id_PROPER_monotonic;assumption.
      intros ;apply P_id_U31_hat_1_monotonic;assumption.
      intros ;apply P_id_U72_hat_1_monotonic;assumption.
      intros ;apply P_id_U61_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNELIST_monotonic;assumption.
      intros ;apply P_id_U41_hat_1_monotonic;assumption.
      intros ;apply P_id_U11_hat_1_monotonic;assumption.
      intros ;apply P_id_U81_hat_1_monotonic;assumption.
      intros ;apply P_id____hat_1_monotonic;assumption.
    Qed.
    
    Ltac rewrite_and_unfold  :=
     do 2 (rewrite marked_measure_equation);
      repeat (
      match goal with
        |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
         rewrite (measure_equation (algebra.Alg.Term f t))
        | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
         rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
        end
      ).
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_21.DP_R_xml_0_scc_21_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_21_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 (x12:Z) := 1* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 3* x12 + 3* x13.
   
   Definition P_id_isList (x12:Z) := 2* x12.
   
   Definition P_id_i  := 1.
   
   Definition P_id_U11 (x12:Z) := 1* x12.
   
   Definition P_id_isQid (x12:Z) := 2* x12.
   
   Definition P_id_isNeList (x12:Z) := 1* x12.
   
   Definition P_id_ok (x12:Z) := 1 + 2* x12.
   
   Definition P_id_mark (x12:Z) := 0.
   
   Definition P_id_isPal (x12:Z) := 1* x12.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_u  := 3.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_a  := 1.
   
   Definition P_id_U52 (x12:Z) := 3* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_U72 (x12:Z) := 2* x12.
   
   Definition P_id_U31 (x12:Z) := 2* x12.
   
   Definition P_id_o  := 1.
   
   Definition P_id_tt  := 3.
   
   Definition P_id_isNePal (x12:Z) := 2* x12.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 2.
   
   Definition P_id_U81 (x12:Z) := 1* x12.
   
   Definition P_id_U42 (x12:Z) := 2* x12.
   
   Definition P_id_proper (x12:Z) := 3* x12.
   
   Definition P_id_U22 (x12:Z) := 1* x12.
   
   Definition P_id_e  := 1.
   
   Definition P_id_U61 (x12:Z) := 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
   
   Definition le a b := marked_measure a <= marked_measure b.
   
   Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
   Proof.
     unfold lt, le in *.
     intros a b c.
     apply (interp.le_lt_compat_right (interp.o_Z 0)).
   Qed.
   
   Lemma wf_lt : well_founded lt.
   Proof.
     unfold lt in *.
     apply Inverse_Image.wf_inverse_image with  (B:=Z).
     apply Zwf.Zwf_well_founded.
   Qed.
   
   Lemma DP_R_xml_0_scc_21_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_21_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_21_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_21_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_21_large  := WF_DP_R_xml_0_scc_21_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_21.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_21_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_21_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_21_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_21_large.
   Qed.
  End WF_DP_R_xml_0_scc_21.
  
  Definition wf_DP_R_xml_0_scc_21  := WF_DP_R_xml_0_scc_21.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_21 :
   forall x y, (DP_R_xml_0_scc_21 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_21).
    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_21.
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_22  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(U51(X1_,X2_)),U51(proper(X1_),proper(X2_))> *)
    | DP_R_xml_0_non_scc_22_0 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U51 (x9::x10::nil)) x12) ->
       DP_R_xml_0_non_scc_22 (algebra.Alg.Term algebra.F.id_U51 
                              ((algebra.Alg.Term algebra.F.id_proper 
                              (x9::nil))::(algebra.Alg.Term 
                              algebra.F.id_proper (x10::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::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_21;
        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_23  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <U42(ok(X_)),U42(X_)> *)
    | DP_R_xml_0_scc_23_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
       DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_U42 (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_U42 (x12::nil))
     (* <U42(mark(X_)),U42(X_)> *)
    | DP_R_xml_0_scc_23_1 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
       DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_U42 (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_U42 (x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_23.
   Inductive DP_R_xml_0_scc_23_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U42(ok(X_)),U42(X_)> *)
     | DP_R_xml_0_scc_23_large_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
        DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_U42 (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_U42 (x12::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_23_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U42(mark(X_)),U42(X_)> *)
     | DP_R_xml_0_scc_23_strict_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
        DP_R_xml_0_scc_23_strict (algebra.Alg.Term algebra.F.id_U42 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_U42 (x12::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_23_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 (x12:Z) := 2* x12.
    
    Definition P_id_U71 (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_isList (x12:Z) := 2 + 2* x12.
    
    Definition P_id_i  := 2.
    
    Definition P_id_U11 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_isQid (x12:Z) := 2 + 2* x12.
    
    Definition P_id_isNeList (x12:Z) := 2* x12.
    
    Definition P_id_ok (x12:Z) := 1 + 1* x12.
    
    Definition P_id_mark (x12:Z) := 0.
    
    Definition P_id_isPal (x12:Z) := 1* x12.
    
    Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
    
    Definition P_id_u  := 1.
    
    Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_a  := 1.
    
    Definition P_id_U52 (x12:Z) := 2* x12.
    
    Definition P_id___ (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_U72 (x12:Z) := 1* x12.
    
    Definition P_id_U31 (x12:Z) := 2 + 2* x12.
    
    Definition P_id_o  := 1.
    
    Definition P_id_tt  := 2.
    
    Definition P_id_isNePal (x12:Z) := 2* x12.
    
    Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
    
    Definition P_id_top (x12:Z) := 0.
    
    Definition P_id_nil  := 2.
    
    Definition P_id_U81 (x12:Z) := 2* x12.
    
    Definition P_id_U42 (x12:Z) := 1 + 2* x12.
    
    Definition P_id_proper (x12:Z) := 2* x12.
    
    Definition P_id_U22 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_e  := 2.
    
    Definition P_id_U61 (x12:Z) := 1* x12.
    
    Lemma P_id_active_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_isList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
    Proof.
      intros x13 x12.
      intros [H_1 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 x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U72_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_top_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
    Proof.
      intros x13 x12.
      intros [H_1 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_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_active x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_i_bounded : 0 <= P_id_i .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
    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 x12, (0 <= x12) ->0 <= P_id_mark x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_u_bounded : 0 <= P_id_u .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_a_bounded : 0 <= P_id_a .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_o_bounded : 0 <= P_id_o .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_tt_bounded : 0 <= P_id_tt .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
    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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_proper x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_e_bounded : 0 <= P_id_e .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
    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_U71 P_id_isList P_id_i P_id_U11 
       P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
       P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
       P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
       P_id_proper P_id_U22 P_id_e P_id_U61.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                     P_id_active (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                     P_id_U71 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                     P_id_isList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                    | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                     P_id_U11 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                     P_id_isQid (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                     P_id_isNeList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                     P_id_ok (measure x12)
                    | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                     P_id_mark (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                     P_id_isPal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                     P_id_U41 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                    | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                     P_id_U21 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                    | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                     P_id_U52 (measure x12)
                    | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                     P_id___ (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                     P_id_U72 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                     P_id_U31 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                    | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                    | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                     P_id_isNePal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                     P_id_U51 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                     P_id_top (measure x12)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                     P_id_U81 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                     P_id_U42 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                     P_id_proper (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                     P_id_U22 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                    | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                     P_id_U61 (measure x12)
                    | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_U22_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNEPAL (x12:Z) := 0.
    
    Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U52_hat_1 (x12:Z) := 0.
    
    Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U42_hat_1 (x12:Z) := 1* x12.
    
    Definition P_id_TOP (x12:Z) := 0.
    
    Definition P_id_ISQID (x12:Z) := 0.
    
    Definition P_id_ISPAL (x12:Z) := 0.
    
    Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_ACTIVE (x12:Z) := 0.
    
    Definition P_id_ISLIST (x12:Z) := 0.
    
    Definition P_id_PROPER (x12:Z) := 0.
    
    Definition P_id_U31_hat_1 (x12:Z) := 0.
    
    Definition P_id_U72_hat_1 (x12:Z) := 0.
    
    Definition P_id_U61_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNELIST (x12:Z) := 0.
    
    Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U11_hat_1 (x12:Z) := 0.
    
    Definition P_id_U81_hat_1 (x12:Z) := 0.
    
    Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
    
    Lemma P_id_U22_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNEPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_TOP_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISQID_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISLIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_PROPER_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNELIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U11_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Definition marked_measure  := 
      InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
       P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
       P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
       P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
       P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
       P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
       P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
       P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
       P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
       P_id_U81_hat_1 P_id____hat_1.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                            P_id_U22_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNePal 
                              (x12::nil)) =>
                            P_id_ISNEPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U21 (x13::
                              x12::nil)) =>
                            P_id_U21_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                            P_id_U52_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U51 (x13::
                              x12::nil)) =>
                            P_id_U51_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                            P_id_U42_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                            P_id_TOP (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                            P_id_ISQID (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                            P_id_ISPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U71 (x13::
                              x12::nil)) =>
                            P_id_U71_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_active 
                              (x12::nil)) =>
                            P_id_ACTIVE (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isList 
                              (x12::nil)) =>
                            P_id_ISLIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_proper 
                              (x12::nil)) =>
                            P_id_PROPER (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                            P_id_U31_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                            P_id_U72_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                            P_id_U61_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNeList 
                              (x12::nil)) =>
                            P_id_ISNELIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U41 (x13::
                              x12::nil)) =>
                            P_id_U41_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                            P_id_U11_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                            P_id_U81_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id___ (x13::
                              x12::nil)) =>
                            P_id____hat_1 (measure x13) (measure x12)
                           | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_U22_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNEPAL_monotonic;assumption.
      intros ;apply P_id_U21_hat_1_monotonic;assumption.
      intros ;apply P_id_U52_hat_1_monotonic;assumption.
      intros ;apply P_id_U51_hat_1_monotonic;assumption.
      intros ;apply P_id_U42_hat_1_monotonic;assumption.
      intros ;apply P_id_TOP_monotonic;assumption.
      intros ;apply P_id_ISQID_monotonic;assumption.
      intros ;apply P_id_ISPAL_monotonic;assumption.
      intros ;apply P_id_U71_hat_1_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_ISLIST_monotonic;assumption.
      intros ;apply P_id_PROPER_monotonic;assumption.
      intros ;apply P_id_U31_hat_1_monotonic;assumption.
      intros ;apply P_id_U72_hat_1_monotonic;assumption.
      intros ;apply P_id_U61_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNELIST_monotonic;assumption.
      intros ;apply P_id_U41_hat_1_monotonic;assumption.
      intros ;apply P_id_U11_hat_1_monotonic;assumption.
      intros ;apply P_id_U81_hat_1_monotonic;assumption.
      intros ;apply P_id____hat_1_monotonic;assumption.
    Qed.
    
    Ltac rewrite_and_unfold  :=
     do 2 (rewrite marked_measure_equation);
      repeat (
      match goal with
        |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
         rewrite (measure_equation (algebra.Alg.Term f t))
        | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
         rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
        end
      ).
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_23.DP_R_xml_0_scc_23_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_23_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 (x12:Z) := 1 + 2* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 2* x12.
   
   Definition P_id_isList (x12:Z) := 0.
   
   Definition P_id_i  := 0.
   
   Definition P_id_U11 (x12:Z) := 1* x12.
   
   Definition P_id_isQid (x12:Z) := 0.
   
   Definition P_id_isNeList (x12:Z) := 0.
   
   Definition P_id_ok (x12:Z) := 1* x12.
   
   Definition P_id_mark (x12:Z) := 1 + 1* x12.
   
   Definition P_id_isPal (x12:Z) := 1.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_u  := 0.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_a  := 0.
   
   Definition P_id_U52 (x12:Z) := 1* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 2 + 1* x12 + 2* x13.
   
   Definition P_id_U72 (x12:Z) := 1 + 1* x12.
   
   Definition P_id_U31 (x12:Z) := 1* x12.
   
   Definition P_id_o  := 2.
   
   Definition P_id_tt  := 0.
   
   Definition P_id_isNePal (x12:Z) := 1.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_U81 (x12:Z) := 1* x12.
   
   Definition P_id_U42 (x12:Z) := 1* x12.
   
   Definition P_id_proper (x12:Z) := 2* x12.
   
   Definition P_id_U22 (x12:Z) := 1* x12.
   
   Definition P_id_e  := 0.
   
   Definition P_id_U61 (x12:Z) := 2 + 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 2* x12.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
   
   Definition le a b := marked_measure a <= marked_measure b.
   
   Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
   Proof.
     unfold lt, le in *.
     intros a b c.
     apply (interp.le_lt_compat_right (interp.o_Z 0)).
   Qed.
   
   Lemma wf_lt : well_founded lt.
   Proof.
     unfold lt in *.
     apply Inverse_Image.wf_inverse_image with  (B:=Z).
     apply Zwf.Zwf_well_founded.
   Qed.
   
   Lemma DP_R_xml_0_scc_23_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_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_23_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_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_23_large  := WF_DP_R_xml_0_scc_23_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_23.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_23_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_23_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_23_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_23_large.
   Qed.
  End WF_DP_R_xml_0_scc_23.
  
  Definition wf_DP_R_xml_0_scc_23  := WF_DP_R_xml_0_scc_23.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_23 :
   forall x y, (DP_R_xml_0_scc_23 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_23).
    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_23.
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_24  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(U42(X_)),U42(proper(X_))> *)
    | DP_R_xml_0_non_scc_24_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U42 (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_24 (algebra.Alg.Term algebra.F.id_U42 
                              ((algebra.Alg.Term algebra.F.id_proper 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::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_23;
        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_25  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <U41(ok(X1_),ok(X2_)),U41(X1_,X2_)> *)
    | DP_R_xml_0_scc_25_0 :
     forall x12 x10 x9 x13, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x9::nil)) x13) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x10::nil)) x12) ->
        DP_R_xml_0_scc_25 (algebra.Alg.Term algebra.F.id_U41 (x9::x10::nil)) 
         (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil))
     (* <U41(mark(X1_),X2_),U41(X1_,X2_)> *)
    | DP_R_xml_0_scc_25_1 :
     forall x12 x10 x9 x13, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x9::nil)) x13) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x10 x12) ->
        DP_R_xml_0_scc_25 (algebra.Alg.Term algebra.F.id_U41 (x9::x10::nil)) 
         (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_25.
   Inductive DP_R_xml_0_scc_25_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U41(mark(X1_),X2_),U41(X1_,X2_)> *)
     | DP_R_xml_0_scc_25_large_0 :
      forall x12 x10 x9 x13, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x9::nil)) x13) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   x10 x12) ->
         DP_R_xml_0_scc_25_large (algebra.Alg.Term algebra.F.id_U41 (x9::
                                  x10::nil)) 
          (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_25_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U41(ok(X1_),ok(X2_)),U41(X1_,X2_)> *)
     | DP_R_xml_0_scc_25_strict_0 :
      forall x12 x10 x9 x13, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x9::nil)) x13) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_ok (x10::nil)) x12) ->
         DP_R_xml_0_scc_25_strict (algebra.Alg.Term algebra.F.id_U41 (x9::
                                   x10::nil)) 
          (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_25_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 (x12:Z) := 2* x12.
    
    Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 3* x12 + 1* x13.
    
    Definition P_id_isList (x12:Z) := 2 + 2* x12.
    
    Definition P_id_i  := 2.
    
    Definition P_id_U11 (x12:Z) := 1 + 1* x12.
    
    Definition P_id_isQid (x12:Z) := 2* x12.
    
    Definition P_id_isNeList (x12:Z) := 2 + 1* x12.
    
    Definition P_id_ok (x12:Z) := 0.
    
    Definition P_id_mark (x12:Z) := 1 + 1* x12.
    
    Definition P_id_isPal (x12:Z) := 3 + 2* x12.
    
    Definition P_id_U41 (x12:Z) (x13:Z) := 1 + 2* x12 + 2* x13.
    
    Definition P_id_u  := 2.
    
    Definition P_id_U21 (x12:Z) (x13:Z) := 2* x12 + 3* x13.
    
    Definition P_id_a  := 2.
    
    Definition P_id_U52 (x12:Z) := 1 + 3* x12.
    
    Definition P_id___ (x12:Z) (x13:Z) := 1 + 3* x12 + 2* x13.
    
    Definition P_id_U72 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_U31 (x12:Z) := 1* x12.
    
    Definition P_id_o  := 2.
    
    Definition P_id_tt  := 2.
    
    Definition P_id_isNePal (x12:Z) := 2 + 2* x12.
    
    Definition P_id_U51 (x12:Z) (x13:Z) := 2* x12 + 3* x13.
    
    Definition P_id_top (x12:Z) := 0.
    
    Definition P_id_nil  := 0.
    
    Definition P_id_U81 (x12:Z) := 2* x12.
    
    Definition P_id_U42 (x12:Z) := 1 + 3* x12.
    
    Definition P_id_proper (x12:Z) := 1* x12.
    
    Definition P_id_U22 (x12:Z) := 1 + 2* x12.
    
    Definition P_id_e  := 2.
    
    Definition P_id_U61 (x12:Z) := 2* x12.
    
    Lemma P_id_active_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_isList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
    Proof.
      intros x13 x12.
      intros [H_1 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 x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U72_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_top_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
    Proof.
      intros x13 x12.
      intros [H_1 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_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_active x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_i_bounded : 0 <= P_id_i .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
    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 x12, (0 <= x12) ->0 <= P_id_mark x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_u_bounded : 0 <= P_id_u .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_a_bounded : 0 <= P_id_a .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_o_bounded : 0 <= P_id_o .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_tt_bounded : 0 <= P_id_tt .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
    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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_proper x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_e_bounded : 0 <= P_id_e .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
    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_U71 P_id_isList P_id_i P_id_U11 
       P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
       P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
       P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
       P_id_proper P_id_U22 P_id_e P_id_U61.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                     P_id_active (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                     P_id_U71 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                     P_id_isList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                    | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                     P_id_U11 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                     P_id_isQid (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                     P_id_isNeList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                     P_id_ok (measure x12)
                    | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                     P_id_mark (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                     P_id_isPal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                     P_id_U41 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                    | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                     P_id_U21 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                    | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                     P_id_U52 (measure x12)
                    | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                     P_id___ (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                     P_id_U72 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                     P_id_U31 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                    | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                    | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                     P_id_isNePal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                     P_id_U51 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                     P_id_top (measure x12)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                     P_id_U81 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                     P_id_U42 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                     P_id_proper (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                     P_id_U22 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                    | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                     P_id_U61 (measure x12)
                    | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_U22_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNEPAL (x12:Z) := 0.
    
    Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U52_hat_1 (x12:Z) := 0.
    
    Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U42_hat_1 (x12:Z) := 0.
    
    Definition P_id_TOP (x12:Z) := 0.
    
    Definition P_id_ISQID (x12:Z) := 0.
    
    Definition P_id_ISPAL (x12:Z) := 0.
    
    Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_ACTIVE (x12:Z) := 0.
    
    Definition P_id_ISLIST (x12:Z) := 0.
    
    Definition P_id_PROPER (x12:Z) := 0.
    
    Definition P_id_U31_hat_1 (x12:Z) := 0.
    
    Definition P_id_U72_hat_1 (x12:Z) := 0.
    
    Definition P_id_U61_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNELIST (x12:Z) := 0.
    
    Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 1* x12.
    
    Definition P_id_U11_hat_1 (x12:Z) := 0.
    
    Definition P_id_U81_hat_1 (x12:Z) := 0.
    
    Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
    
    Lemma P_id_U22_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNEPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_TOP_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISQID_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISLIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_PROPER_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNELIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U11_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Definition marked_measure  := 
      InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
       P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
       P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
       P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
       P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
       P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
       P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
       P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
       P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
       P_id_U81_hat_1 P_id____hat_1.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                            P_id_U22_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNePal 
                              (x12::nil)) =>
                            P_id_ISNEPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U21 (x13::
                              x12::nil)) =>
                            P_id_U21_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                            P_id_U52_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U51 (x13::
                              x12::nil)) =>
                            P_id_U51_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                            P_id_U42_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                            P_id_TOP (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                            P_id_ISQID (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                            P_id_ISPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U71 (x13::
                              x12::nil)) =>
                            P_id_U71_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_active 
                              (x12::nil)) =>
                            P_id_ACTIVE (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isList 
                              (x12::nil)) =>
                            P_id_ISLIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_proper 
                              (x12::nil)) =>
                            P_id_PROPER (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                            P_id_U31_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                            P_id_U72_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                            P_id_U61_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNeList 
                              (x12::nil)) =>
                            P_id_ISNELIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U41 (x13::
                              x12::nil)) =>
                            P_id_U41_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                            P_id_U11_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                            P_id_U81_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id___ (x13::
                              x12::nil)) =>
                            P_id____hat_1 (measure x13) (measure x12)
                           | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_U22_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNEPAL_monotonic;assumption.
      intros ;apply P_id_U21_hat_1_monotonic;assumption.
      intros ;apply P_id_U52_hat_1_monotonic;assumption.
      intros ;apply P_id_U51_hat_1_monotonic;assumption.
      intros ;apply P_id_U42_hat_1_monotonic;assumption.
      intros ;apply P_id_TOP_monotonic;assumption.
      intros ;apply P_id_ISQID_monotonic;assumption.
      intros ;apply P_id_ISPAL_monotonic;assumption.
      intros ;apply P_id_U71_hat_1_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_ISLIST_monotonic;assumption.
      intros ;apply P_id_PROPER_monotonic;assumption.
      intros ;apply P_id_U31_hat_1_monotonic;assumption.
      intros ;apply P_id_U72_hat_1_monotonic;assumption.
      intros ;apply P_id_U61_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNELIST_monotonic;assumption.
      intros ;apply P_id_U41_hat_1_monotonic;assumption.
      intros ;apply P_id_U11_hat_1_monotonic;assumption.
      intros ;apply P_id_U81_hat_1_monotonic;assumption.
      intros ;apply P_id____hat_1_monotonic;assumption.
    Qed.
    
    Ltac rewrite_and_unfold  :=
     do 2 (rewrite marked_measure_equation);
      repeat (
      match goal with
        |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
         rewrite (measure_equation (algebra.Alg.Term f t))
        | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
         rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
        end
      ).
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_25.DP_R_xml_0_scc_25_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_25_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 (x12:Z) := 1* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 3* x12 + 3* x13.
   
   Definition P_id_isList (x12:Z) := 2* x12.
   
   Definition P_id_i  := 1.
   
   Definition P_id_U11 (x12:Z) := 1* x12.
   
   Definition P_id_isQid (x12:Z) := 2* x12.
   
   Definition P_id_isNeList (x12:Z) := 1* x12.
   
   Definition P_id_ok (x12:Z) := 1 + 2* x12.
   
   Definition P_id_mark (x12:Z) := 0.
   
   Definition P_id_isPal (x12:Z) := 1* x12.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_u  := 3.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_a  := 1.
   
   Definition P_id_U52 (x12:Z) := 3* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_U72 (x12:Z) := 2* x12.
   
   Definition P_id_U31 (x12:Z) := 2* x12.
   
   Definition P_id_o  := 1.
   
   Definition P_id_tt  := 3.
   
   Definition P_id_isNePal (x12:Z) := 2* x12.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 2.
   
   Definition P_id_U81 (x12:Z) := 1* x12.
   
   Definition P_id_U42 (x12:Z) := 2* x12.
   
   Definition P_id_proper (x12:Z) := 3* x12.
   
   Definition P_id_U22 (x12:Z) := 1* x12.
   
   Definition P_id_e  := 1.
   
   Definition P_id_U61 (x12:Z) := 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
   
   Definition le a b := marked_measure a <= marked_measure b.
   
   Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
   Proof.
     unfold lt, le in *.
     intros a b c.
     apply (interp.le_lt_compat_right (interp.o_Z 0)).
   Qed.
   
   Lemma wf_lt : well_founded lt.
   Proof.
     unfold lt in *.
     apply Inverse_Image.wf_inverse_image with  (B:=Z).
     apply Zwf.Zwf_well_founded.
   Qed.
   
   Lemma DP_R_xml_0_scc_25_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_25_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_25_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_25_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_25_large  := WF_DP_R_xml_0_scc_25_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_25.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_25_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_25_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_25_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_25_large.
   Qed.
  End WF_DP_R_xml_0_scc_25.
  
  Definition wf_DP_R_xml_0_scc_25  := WF_DP_R_xml_0_scc_25.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_25 :
   forall x y, (DP_R_xml_0_scc_25 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_25).
    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_25.
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_26  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(U41(X1_,X2_)),U41(proper(X1_),proper(X2_))> *)
    | DP_R_xml_0_non_scc_26_0 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U41 (x9::x10::nil)) x12) ->
       DP_R_xml_0_non_scc_26 (algebra.Alg.Term algebra.F.id_U41 
                              ((algebra.Alg.Term algebra.F.id_proper 
                              (x9::nil))::(algebra.Alg.Term 
                              algebra.F.id_proper (x10::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::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_25;
        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_27  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <U31(ok(X_)),U31(X_)> *)
    | DP_R_xml_0_scc_27_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
       DP_R_xml_0_scc_27 (algebra.Alg.Term algebra.F.id_U31 (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_U31 (x12::nil))
     (* <U31(mark(X_)),U31(X_)> *)
    | DP_R_xml_0_scc_27_1 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
       DP_R_xml_0_scc_27 (algebra.Alg.Term algebra.F.id_U31 (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_U31 (x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_27.
   Inductive DP_R_xml_0_scc_27_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U31(ok(X_)),U31(X_)> *)
     | DP_R_xml_0_scc_27_large_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
        DP_R_xml_0_scc_27_large (algebra.Alg.Term algebra.F.id_U31 (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_U31 (x12::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_27_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U31(mark(X_)),U31(X_)> *)
     | DP_R_xml_0_scc_27_strict_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
        DP_R_xml_0_scc_27_strict (algebra.Alg.Term algebra.F.id_U31 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_U31 (x12::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_27_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 (x12:Z) := 2* x12.
    
    Definition P_id_U71 (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_isList (x12:Z) := 2 + 2* x12.
    
    Definition P_id_i  := 2.
    
    Definition P_id_U11 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_isQid (x12:Z) := 2 + 2* x12.
    
    Definition P_id_isNeList (x12:Z) := 2* x12.
    
    Definition P_id_ok (x12:Z) := 1 + 1* x12.
    
    Definition P_id_mark (x12:Z) := 0.
    
    Definition P_id_isPal (x12:Z) := 1* x12.
    
    Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
    
    Definition P_id_u  := 1.
    
    Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_a  := 1.
    
    Definition P_id_U52 (x12:Z) := 2* x12.
    
    Definition P_id___ (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_U72 (x12:Z) := 1* x12.
    
    Definition P_id_U31 (x12:Z) := 2 + 2* x12.
    
    Definition P_id_o  := 1.
    
    Definition P_id_tt  := 2.
    
    Definition P_id_isNePal (x12:Z) := 2* x12.
    
    Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
    
    Definition P_id_top (x12:Z) := 0.
    
    Definition P_id_nil  := 2.
    
    Definition P_id_U81 (x12:Z) := 2* x12.
    
    Definition P_id_U42 (x12:Z) := 1 + 2* x12.
    
    Definition P_id_proper (x12:Z) := 2* x12.
    
    Definition P_id_U22 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_e  := 2.
    
    Definition P_id_U61 (x12:Z) := 1* x12.
    
    Lemma P_id_active_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_isList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
    Proof.
      intros x13 x12.
      intros [H_1 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 x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U72_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_top_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
    Proof.
      intros x13 x12.
      intros [H_1 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_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_active x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_i_bounded : 0 <= P_id_i .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
    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 x12, (0 <= x12) ->0 <= P_id_mark x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_u_bounded : 0 <= P_id_u .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_a_bounded : 0 <= P_id_a .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_o_bounded : 0 <= P_id_o .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_tt_bounded : 0 <= P_id_tt .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
    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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_proper x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_e_bounded : 0 <= P_id_e .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
    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_U71 P_id_isList P_id_i P_id_U11 
       P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
       P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
       P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
       P_id_proper P_id_U22 P_id_e P_id_U61.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                     P_id_active (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                     P_id_U71 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                     P_id_isList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                    | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                     P_id_U11 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                     P_id_isQid (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                     P_id_isNeList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                     P_id_ok (measure x12)
                    | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                     P_id_mark (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                     P_id_isPal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                     P_id_U41 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                    | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                     P_id_U21 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                    | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                     P_id_U52 (measure x12)
                    | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                     P_id___ (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                     P_id_U72 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                     P_id_U31 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                    | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                    | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                     P_id_isNePal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                     P_id_U51 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                     P_id_top (measure x12)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                     P_id_U81 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                     P_id_U42 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                     P_id_proper (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                     P_id_U22 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                    | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                     P_id_U61 (measure x12)
                    | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_U22_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNEPAL (x12:Z) := 0.
    
    Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U52_hat_1 (x12:Z) := 0.
    
    Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U42_hat_1 (x12:Z) := 0.
    
    Definition P_id_TOP (x12:Z) := 0.
    
    Definition P_id_ISQID (x12:Z) := 0.
    
    Definition P_id_ISPAL (x12:Z) := 0.
    
    Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_ACTIVE (x12:Z) := 0.
    
    Definition P_id_ISLIST (x12:Z) := 0.
    
    Definition P_id_PROPER (x12:Z) := 0.
    
    Definition P_id_U31_hat_1 (x12:Z) := 1* x12.
    
    Definition P_id_U72_hat_1 (x12:Z) := 0.
    
    Definition P_id_U61_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNELIST (x12:Z) := 0.
    
    Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U11_hat_1 (x12:Z) := 0.
    
    Definition P_id_U81_hat_1 (x12:Z) := 0.
    
    Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
    
    Lemma P_id_U22_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNEPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_TOP_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISQID_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISLIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_PROPER_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNELIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U11_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Definition marked_measure  := 
      InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
       P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
       P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
       P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
       P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
       P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
       P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
       P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
       P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
       P_id_U81_hat_1 P_id____hat_1.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                            P_id_U22_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNePal 
                              (x12::nil)) =>
                            P_id_ISNEPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U21 (x13::
                              x12::nil)) =>
                            P_id_U21_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                            P_id_U52_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U51 (x13::
                              x12::nil)) =>
                            P_id_U51_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                            P_id_U42_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                            P_id_TOP (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                            P_id_ISQID (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                            P_id_ISPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U71 (x13::
                              x12::nil)) =>
                            P_id_U71_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_active 
                              (x12::nil)) =>
                            P_id_ACTIVE (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isList 
                              (x12::nil)) =>
                            P_id_ISLIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_proper 
                              (x12::nil)) =>
                            P_id_PROPER (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                            P_id_U31_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                            P_id_U72_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                            P_id_U61_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNeList 
                              (x12::nil)) =>
                            P_id_ISNELIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U41 (x13::
                              x12::nil)) =>
                            P_id_U41_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                            P_id_U11_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                            P_id_U81_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id___ (x13::
                              x12::nil)) =>
                            P_id____hat_1 (measure x13) (measure x12)
                           | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_U22_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNEPAL_monotonic;assumption.
      intros ;apply P_id_U21_hat_1_monotonic;assumption.
      intros ;apply P_id_U52_hat_1_monotonic;assumption.
      intros ;apply P_id_U51_hat_1_monotonic;assumption.
      intros ;apply P_id_U42_hat_1_monotonic;assumption.
      intros ;apply P_id_TOP_monotonic;assumption.
      intros ;apply P_id_ISQID_monotonic;assumption.
      intros ;apply P_id_ISPAL_monotonic;assumption.
      intros ;apply P_id_U71_hat_1_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_ISLIST_monotonic;assumption.
      intros ;apply P_id_PROPER_monotonic;assumption.
      intros ;apply P_id_U31_hat_1_monotonic;assumption.
      intros ;apply P_id_U72_hat_1_monotonic;assumption.
      intros ;apply P_id_U61_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNELIST_monotonic;assumption.
      intros ;apply P_id_U41_hat_1_monotonic;assumption.
      intros ;apply P_id_U11_hat_1_monotonic;assumption.
      intros ;apply P_id_U81_hat_1_monotonic;assumption.
      intros ;apply P_id____hat_1_monotonic;assumption.
    Qed.
    
    Ltac rewrite_and_unfold  :=
     do 2 (rewrite marked_measure_equation);
      repeat (
      match goal with
        |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
         rewrite (measure_equation (algebra.Alg.Term f t))
        | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
         rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
        end
      ).
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_27.DP_R_xml_0_scc_27_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_27_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 (x12:Z) := 1 + 2* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 2* x12.
   
   Definition P_id_isList (x12:Z) := 0.
   
   Definition P_id_i  := 0.
   
   Definition P_id_U11 (x12:Z) := 1* x12.
   
   Definition P_id_isQid (x12:Z) := 0.
   
   Definition P_id_isNeList (x12:Z) := 0.
   
   Definition P_id_ok (x12:Z) := 1* x12.
   
   Definition P_id_mark (x12:Z) := 1 + 1* x12.
   
   Definition P_id_isPal (x12:Z) := 1.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_u  := 0.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_a  := 0.
   
   Definition P_id_U52 (x12:Z) := 1* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 2 + 1* x12 + 2* x13.
   
   Definition P_id_U72 (x12:Z) := 1 + 1* x12.
   
   Definition P_id_U31 (x12:Z) := 1* x12.
   
   Definition P_id_o  := 2.
   
   Definition P_id_tt  := 0.
   
   Definition P_id_isNePal (x12:Z) := 1.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_U81 (x12:Z) := 1* x12.
   
   Definition P_id_U42 (x12:Z) := 1* x12.
   
   Definition P_id_proper (x12:Z) := 2* x12.
   
   Definition P_id_U22 (x12:Z) := 1* x12.
   
   Definition P_id_e  := 0.
   
   Definition P_id_U61 (x12:Z) := 2 + 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 2* x12.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
   
   Definition le a b := marked_measure a <= marked_measure b.
   
   Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
   Proof.
     unfold lt, le in *.
     intros a b c.
     apply (interp.le_lt_compat_right (interp.o_Z 0)).
   Qed.
   
   Lemma wf_lt : well_founded lt.
   Proof.
     unfold lt in *.
     apply Inverse_Image.wf_inverse_image with  (B:=Z).
     apply Zwf.Zwf_well_founded.
   Qed.
   
   Lemma DP_R_xml_0_scc_27_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_27_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_27_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_27_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_27_large  := WF_DP_R_xml_0_scc_27_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_27.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_27_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_27_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_27_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_27_large.
   Qed.
  End WF_DP_R_xml_0_scc_27.
  
  Definition wf_DP_R_xml_0_scc_27  := WF_DP_R_xml_0_scc_27.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_27 :
   forall x y, (DP_R_xml_0_scc_27 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_27).
    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_27.
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_28  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(U31(X_)),U31(proper(X_))> *)
    | DP_R_xml_0_non_scc_28_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U31 (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_28 (algebra.Alg.Term algebra.F.id_U31 
                              ((algebra.Alg.Term algebra.F.id_proper 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::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_27;
        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_29  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <U22(ok(X_)),U22(X_)> *)
    | DP_R_xml_0_scc_29_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
       DP_R_xml_0_scc_29 (algebra.Alg.Term algebra.F.id_U22 (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_U22 (x12::nil))
     (* <U22(mark(X_)),U22(X_)> *)
    | DP_R_xml_0_scc_29_1 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
       DP_R_xml_0_scc_29 (algebra.Alg.Term algebra.F.id_U22 (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_U22 (x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_29.
   Inductive DP_R_xml_0_scc_29_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U22(ok(X_)),U22(X_)> *)
     | DP_R_xml_0_scc_29_large_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
        DP_R_xml_0_scc_29_large (algebra.Alg.Term algebra.F.id_U22 (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_U22 (x12::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_29_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U22(mark(X_)),U22(X_)> *)
     | DP_R_xml_0_scc_29_strict_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
        DP_R_xml_0_scc_29_strict (algebra.Alg.Term algebra.F.id_U22 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_U22 (x12::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_29_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 (x12:Z) := 2* x12.
    
    Definition P_id_U71 (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_isList (x12:Z) := 2 + 2* x12.
    
    Definition P_id_i  := 2.
    
    Definition P_id_U11 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_isQid (x12:Z) := 2 + 2* x12.
    
    Definition P_id_isNeList (x12:Z) := 2* x12.
    
    Definition P_id_ok (x12:Z) := 1 + 1* x12.
    
    Definition P_id_mark (x12:Z) := 0.
    
    Definition P_id_isPal (x12:Z) := 1* x12.
    
    Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
    
    Definition P_id_u  := 1.
    
    Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_a  := 1.
    
    Definition P_id_U52 (x12:Z) := 2* x12.
    
    Definition P_id___ (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_U72 (x12:Z) := 1* x12.
    
    Definition P_id_U31 (x12:Z) := 2 + 2* x12.
    
    Definition P_id_o  := 1.
    
    Definition P_id_tt  := 2.
    
    Definition P_id_isNePal (x12:Z) := 2* x12.
    
    Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
    
    Definition P_id_top (x12:Z) := 0.
    
    Definition P_id_nil  := 2.
    
    Definition P_id_U81 (x12:Z) := 2* x12.
    
    Definition P_id_U42 (x12:Z) := 1 + 2* x12.
    
    Definition P_id_proper (x12:Z) := 2* x12.
    
    Definition P_id_U22 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_e  := 2.
    
    Definition P_id_U61 (x12:Z) := 1* x12.
    
    Lemma P_id_active_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_isList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
    Proof.
      intros x13 x12.
      intros [H_1 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 x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U72_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_top_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
    Proof.
      intros x13 x12.
      intros [H_1 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_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_active x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_i_bounded : 0 <= P_id_i .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
    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 x12, (0 <= x12) ->0 <= P_id_mark x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_u_bounded : 0 <= P_id_u .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_a_bounded : 0 <= P_id_a .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_o_bounded : 0 <= P_id_o .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_tt_bounded : 0 <= P_id_tt .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
    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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_proper x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_e_bounded : 0 <= P_id_e .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
    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_U71 P_id_isList P_id_i P_id_U11 
       P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
       P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
       P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
       P_id_proper P_id_U22 P_id_e P_id_U61.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                     P_id_active (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                     P_id_U71 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                     P_id_isList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                    | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                     P_id_U11 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                     P_id_isQid (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                     P_id_isNeList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                     P_id_ok (measure x12)
                    | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                     P_id_mark (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                     P_id_isPal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                     P_id_U41 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                    | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                     P_id_U21 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                    | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                     P_id_U52 (measure x12)
                    | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                     P_id___ (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                     P_id_U72 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                     P_id_U31 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                    | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                    | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                     P_id_isNePal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                     P_id_U51 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                     P_id_top (measure x12)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                     P_id_U81 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                     P_id_U42 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                     P_id_proper (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                     P_id_U22 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                    | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                     P_id_U61 (measure x12)
                    | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_U22_hat_1 (x12:Z) := 1* x12.
    
    Definition P_id_ISNEPAL (x12:Z) := 0.
    
    Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U52_hat_1 (x12:Z) := 0.
    
    Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U42_hat_1 (x12:Z) := 0.
    
    Definition P_id_TOP (x12:Z) := 0.
    
    Definition P_id_ISQID (x12:Z) := 0.
    
    Definition P_id_ISPAL (x12:Z) := 0.
    
    Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_ACTIVE (x12:Z) := 0.
    
    Definition P_id_ISLIST (x12:Z) := 0.
    
    Definition P_id_PROPER (x12:Z) := 0.
    
    Definition P_id_U31_hat_1 (x12:Z) := 0.
    
    Definition P_id_U72_hat_1 (x12:Z) := 0.
    
    Definition P_id_U61_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNELIST (x12:Z) := 0.
    
    Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U11_hat_1 (x12:Z) := 0.
    
    Definition P_id_U81_hat_1 (x12:Z) := 0.
    
    Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
    
    Lemma P_id_U22_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNEPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_TOP_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISQID_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISLIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_PROPER_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNELIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U11_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Definition marked_measure  := 
      InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
       P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
       P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
       P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
       P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
       P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
       P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
       P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
       P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
       P_id_U81_hat_1 P_id____hat_1.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                            P_id_U22_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNePal 
                              (x12::nil)) =>
                            P_id_ISNEPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U21 (x13::
                              x12::nil)) =>
                            P_id_U21_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                            P_id_U52_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U51 (x13::
                              x12::nil)) =>
                            P_id_U51_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                            P_id_U42_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                            P_id_TOP (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                            P_id_ISQID (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                            P_id_ISPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U71 (x13::
                              x12::nil)) =>
                            P_id_U71_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_active 
                              (x12::nil)) =>
                            P_id_ACTIVE (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isList 
                              (x12::nil)) =>
                            P_id_ISLIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_proper 
                              (x12::nil)) =>
                            P_id_PROPER (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                            P_id_U31_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                            P_id_U72_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                            P_id_U61_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNeList 
                              (x12::nil)) =>
                            P_id_ISNELIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U41 (x13::
                              x12::nil)) =>
                            P_id_U41_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                            P_id_U11_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                            P_id_U81_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id___ (x13::
                              x12::nil)) =>
                            P_id____hat_1 (measure x13) (measure x12)
                           | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_U22_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNEPAL_monotonic;assumption.
      intros ;apply P_id_U21_hat_1_monotonic;assumption.
      intros ;apply P_id_U52_hat_1_monotonic;assumption.
      intros ;apply P_id_U51_hat_1_monotonic;assumption.
      intros ;apply P_id_U42_hat_1_monotonic;assumption.
      intros ;apply P_id_TOP_monotonic;assumption.
      intros ;apply P_id_ISQID_monotonic;assumption.
      intros ;apply P_id_ISPAL_monotonic;assumption.
      intros ;apply P_id_U71_hat_1_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_ISLIST_monotonic;assumption.
      intros ;apply P_id_PROPER_monotonic;assumption.
      intros ;apply P_id_U31_hat_1_monotonic;assumption.
      intros ;apply P_id_U72_hat_1_monotonic;assumption.
      intros ;apply P_id_U61_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNELIST_monotonic;assumption.
      intros ;apply P_id_U41_hat_1_monotonic;assumption.
      intros ;apply P_id_U11_hat_1_monotonic;assumption.
      intros ;apply P_id_U81_hat_1_monotonic;assumption.
      intros ;apply P_id____hat_1_monotonic;assumption.
    Qed.
    
    Ltac rewrite_and_unfold  :=
     do 2 (rewrite marked_measure_equation);
      repeat (
      match goal with
        |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
         rewrite (measure_equation (algebra.Alg.Term f t))
        | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
         rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
        end
      ).
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_29.DP_R_xml_0_scc_29_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_29_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 (x12:Z) := 1 + 2* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 2* x12.
   
   Definition P_id_isList (x12:Z) := 0.
   
   Definition P_id_i  := 0.
   
   Definition P_id_U11 (x12:Z) := 1* x12.
   
   Definition P_id_isQid (x12:Z) := 0.
   
   Definition P_id_isNeList (x12:Z) := 0.
   
   Definition P_id_ok (x12:Z) := 1* x12.
   
   Definition P_id_mark (x12:Z) := 1 + 1* x12.
   
   Definition P_id_isPal (x12:Z) := 1.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_u  := 0.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_a  := 0.
   
   Definition P_id_U52 (x12:Z) := 1* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 2 + 1* x12 + 2* x13.
   
   Definition P_id_U72 (x12:Z) := 1 + 1* x12.
   
   Definition P_id_U31 (x12:Z) := 1* x12.
   
   Definition P_id_o  := 2.
   
   Definition P_id_tt  := 0.
   
   Definition P_id_isNePal (x12:Z) := 1.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_U81 (x12:Z) := 1* x12.
   
   Definition P_id_U42 (x12:Z) := 1* x12.
   
   Definition P_id_proper (x12:Z) := 2* x12.
   
   Definition P_id_U22 (x12:Z) := 1* x12.
   
   Definition P_id_e  := 0.
   
   Definition P_id_U61 (x12:Z) := 2 + 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 2* x12.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
   
   Definition le a b := marked_measure a <= marked_measure b.
   
   Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
   Proof.
     unfold lt, le in *.
     intros a b c.
     apply (interp.le_lt_compat_right (interp.o_Z 0)).
   Qed.
   
   Lemma wf_lt : well_founded lt.
   Proof.
     unfold lt in *.
     apply Inverse_Image.wf_inverse_image with  (B:=Z).
     apply Zwf.Zwf_well_founded.
   Qed.
   
   Lemma DP_R_xml_0_scc_29_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_29_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_29_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_29_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_29_large  := WF_DP_R_xml_0_scc_29_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_29.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_29_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_29_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_29_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_29_large.
   Qed.
  End WF_DP_R_xml_0_scc_29.
  
  Definition wf_DP_R_xml_0_scc_29  := WF_DP_R_xml_0_scc_29.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_29 :
   forall x y, (DP_R_xml_0_scc_29 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_29).
    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_29.
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_30  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(U22(X_)),U22(proper(X_))> *)
    | DP_R_xml_0_non_scc_30_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U22 (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_30 (algebra.Alg.Term algebra.F.id_U22 
                              ((algebra.Alg.Term algebra.F.id_proper 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::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_29;
        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_31  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <U21(ok(X1_),ok(X2_)),U21(X1_,X2_)> *)
    | DP_R_xml_0_scc_31_0 :
     forall x12 x10 x9 x13, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x9::nil)) x13) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x10::nil)) x12) ->
        DP_R_xml_0_scc_31 (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) 
         (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil))
     (* <U21(mark(X1_),X2_),U21(X1_,X2_)> *)
    | DP_R_xml_0_scc_31_1 :
     forall x12 x10 x9 x13, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x9::nil)) x13) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x10 x12) ->
        DP_R_xml_0_scc_31 (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) 
         (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_31.
   Inductive DP_R_xml_0_scc_31_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U21(mark(X1_),X2_),U21(X1_,X2_)> *)
     | DP_R_xml_0_scc_31_large_0 :
      forall x12 x10 x9 x13, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x9::nil)) x13) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   x10 x12) ->
         DP_R_xml_0_scc_31_large (algebra.Alg.Term algebra.F.id_U21 (x9::
                                  x10::nil)) 
          (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_31_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U21(ok(X1_),ok(X2_)),U21(X1_,X2_)> *)
     | DP_R_xml_0_scc_31_strict_0 :
      forall x12 x10 x9 x13, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x9::nil)) x13) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_ok (x10::nil)) x12) ->
         DP_R_xml_0_scc_31_strict (algebra.Alg.Term algebra.F.id_U21 (x9::
                                   x10::nil)) 
          (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_31_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 (x12:Z) := 2* x12.
    
    Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 3* x12 + 1* x13.
    
    Definition P_id_isList (x12:Z) := 2 + 2* x12.
    
    Definition P_id_i  := 2.
    
    Definition P_id_U11 (x12:Z) := 1 + 1* x12.
    
    Definition P_id_isQid (x12:Z) := 2* x12.
    
    Definition P_id_isNeList (x12:Z) := 2 + 1* x12.
    
    Definition P_id_ok (x12:Z) := 0.
    
    Definition P_id_mark (x12:Z) := 1 + 1* x12.
    
    Definition P_id_isPal (x12:Z) := 3 + 2* x12.
    
    Definition P_id_U41 (x12:Z) (x13:Z) := 1 + 2* x12 + 2* x13.
    
    Definition P_id_u  := 2.
    
    Definition P_id_U21 (x12:Z) (x13:Z) := 2* x12 + 3* x13.
    
    Definition P_id_a  := 2.
    
    Definition P_id_U52 (x12:Z) := 1 + 3* x12.
    
    Definition P_id___ (x12:Z) (x13:Z) := 1 + 3* x12 + 2* x13.
    
    Definition P_id_U72 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_U31 (x12:Z) := 1* x12.
    
    Definition P_id_o  := 2.
    
    Definition P_id_tt  := 2.
    
    Definition P_id_isNePal (x12:Z) := 2 + 2* x12.
    
    Definition P_id_U51 (x12:Z) (x13:Z) := 2* x12 + 3* x13.
    
    Definition P_id_top (x12:Z) := 0.
    
    Definition P_id_nil  := 0.
    
    Definition P_id_U81 (x12:Z) := 2* x12.
    
    Definition P_id_U42 (x12:Z) := 1 + 3* x12.
    
    Definition P_id_proper (x12:Z) := 1* x12.
    
    Definition P_id_U22 (x12:Z) := 1 + 2* x12.
    
    Definition P_id_e  := 2.
    
    Definition P_id_U61 (x12:Z) := 2* x12.
    
    Lemma P_id_active_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_isList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
    Proof.
      intros x13 x12.
      intros [H_1 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 x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U72_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_top_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
    Proof.
      intros x13 x12.
      intros [H_1 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_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_active x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_i_bounded : 0 <= P_id_i .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
    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 x12, (0 <= x12) ->0 <= P_id_mark x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_u_bounded : 0 <= P_id_u .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_a_bounded : 0 <= P_id_a .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_o_bounded : 0 <= P_id_o .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_tt_bounded : 0 <= P_id_tt .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
    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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_proper x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_e_bounded : 0 <= P_id_e .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
    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_U71 P_id_isList P_id_i P_id_U11 
       P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
       P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
       P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
       P_id_proper P_id_U22 P_id_e P_id_U61.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                     P_id_active (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                     P_id_U71 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                     P_id_isList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                    | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                     P_id_U11 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                     P_id_isQid (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                     P_id_isNeList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                     P_id_ok (measure x12)
                    | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                     P_id_mark (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                     P_id_isPal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                     P_id_U41 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                    | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                     P_id_U21 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                    | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                     P_id_U52 (measure x12)
                    | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                     P_id___ (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                     P_id_U72 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                     P_id_U31 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                    | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                    | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                     P_id_isNePal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                     P_id_U51 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                     P_id_top (measure x12)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                     P_id_U81 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                     P_id_U42 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                     P_id_proper (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                     P_id_U22 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                    | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                     P_id_U61 (measure x12)
                    | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_U22_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNEPAL (x12:Z) := 0.
    
    Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 1* x12.
    
    Definition P_id_U52_hat_1 (x12:Z) := 0.
    
    Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U42_hat_1 (x12:Z) := 0.
    
    Definition P_id_TOP (x12:Z) := 0.
    
    Definition P_id_ISQID (x12:Z) := 0.
    
    Definition P_id_ISPAL (x12:Z) := 0.
    
    Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_ACTIVE (x12:Z) := 0.
    
    Definition P_id_ISLIST (x12:Z) := 0.
    
    Definition P_id_PROPER (x12:Z) := 0.
    
    Definition P_id_U31_hat_1 (x12:Z) := 0.
    
    Definition P_id_U72_hat_1 (x12:Z) := 0.
    
    Definition P_id_U61_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNELIST (x12:Z) := 0.
    
    Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U11_hat_1 (x12:Z) := 0.
    
    Definition P_id_U81_hat_1 (x12:Z) := 0.
    
    Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
    
    Lemma P_id_U22_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNEPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_TOP_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISQID_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISLIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_PROPER_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNELIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U11_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Definition marked_measure  := 
      InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
       P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
       P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
       P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
       P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
       P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
       P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
       P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
       P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
       P_id_U81_hat_1 P_id____hat_1.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                            P_id_U22_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNePal 
                              (x12::nil)) =>
                            P_id_ISNEPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U21 (x13::
                              x12::nil)) =>
                            P_id_U21_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                            P_id_U52_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U51 (x13::
                              x12::nil)) =>
                            P_id_U51_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                            P_id_U42_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                            P_id_TOP (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                            P_id_ISQID (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                            P_id_ISPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U71 (x13::
                              x12::nil)) =>
                            P_id_U71_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_active 
                              (x12::nil)) =>
                            P_id_ACTIVE (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isList 
                              (x12::nil)) =>
                            P_id_ISLIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_proper 
                              (x12::nil)) =>
                            P_id_PROPER (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                            P_id_U31_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                            P_id_U72_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                            P_id_U61_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNeList 
                              (x12::nil)) =>
                            P_id_ISNELIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U41 (x13::
                              x12::nil)) =>
                            P_id_U41_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                            P_id_U11_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                            P_id_U81_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id___ (x13::
                              x12::nil)) =>
                            P_id____hat_1 (measure x13) (measure x12)
                           | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_U22_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNEPAL_monotonic;assumption.
      intros ;apply P_id_U21_hat_1_monotonic;assumption.
      intros ;apply P_id_U52_hat_1_monotonic;assumption.
      intros ;apply P_id_U51_hat_1_monotonic;assumption.
      intros ;apply P_id_U42_hat_1_monotonic;assumption.
      intros ;apply P_id_TOP_monotonic;assumption.
      intros ;apply P_id_ISQID_monotonic;assumption.
      intros ;apply P_id_ISPAL_monotonic;assumption.
      intros ;apply P_id_U71_hat_1_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_ISLIST_monotonic;assumption.
      intros ;apply P_id_PROPER_monotonic;assumption.
      intros ;apply P_id_U31_hat_1_monotonic;assumption.
      intros ;apply P_id_U72_hat_1_monotonic;assumption.
      intros ;apply P_id_U61_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNELIST_monotonic;assumption.
      intros ;apply P_id_U41_hat_1_monotonic;assumption.
      intros ;apply P_id_U11_hat_1_monotonic;assumption.
      intros ;apply P_id_U81_hat_1_monotonic;assumption.
      intros ;apply P_id____hat_1_monotonic;assumption.
    Qed.
    
    Ltac rewrite_and_unfold  :=
     do 2 (rewrite marked_measure_equation);
      repeat (
      match goal with
        |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
         rewrite (measure_equation (algebra.Alg.Term f t))
        | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
         rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
        end
      ).
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_31.DP_R_xml_0_scc_31_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_31_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 (x12:Z) := 1* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 3* x12 + 3* x13.
   
   Definition P_id_isList (x12:Z) := 2* x12.
   
   Definition P_id_i  := 1.
   
   Definition P_id_U11 (x12:Z) := 1* x12.
   
   Definition P_id_isQid (x12:Z) := 2* x12.
   
   Definition P_id_isNeList (x12:Z) := 1* x12.
   
   Definition P_id_ok (x12:Z) := 1 + 2* x12.
   
   Definition P_id_mark (x12:Z) := 0.
   
   Definition P_id_isPal (x12:Z) := 1* x12.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_u  := 3.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
   
   Definition P_id_a  := 1.
   
   Definition P_id_U52 (x12:Z) := 3* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_U72 (x12:Z) := 2* x12.
   
   Definition P_id_U31 (x12:Z) := 2* x12.
   
   Definition P_id_o  := 1.
   
   Definition P_id_tt  := 3.
   
   Definition P_id_isNePal (x12:Z) := 2* x12.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 2.
   
   Definition P_id_U81 (x12:Z) := 1* x12.
   
   Definition P_id_U42 (x12:Z) := 2* x12.
   
   Definition P_id_proper (x12:Z) := 3* x12.
   
   Definition P_id_U22 (x12:Z) := 1* x12.
   
   Definition P_id_e  := 1.
   
   Definition P_id_U61 (x12:Z) := 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 1* x13.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
   
   Definition le a b := marked_measure a <= marked_measure b.
   
   Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
   Proof.
     unfold lt, le in *.
     intros a b c.
     apply (interp.le_lt_compat_right (interp.o_Z 0)).
   Qed.
   
   Lemma wf_lt : well_founded lt.
   Proof.
     unfold lt in *.
     apply Inverse_Image.wf_inverse_image with  (B:=Z).
     apply Zwf.Zwf_well_founded.
   Qed.
   
   Lemma DP_R_xml_0_scc_31_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_31_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_31_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_31_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_31_large  := WF_DP_R_xml_0_scc_31_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_31.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_31_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_31_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_31_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_31_large.
   Qed.
  End WF_DP_R_xml_0_scc_31.
  
  Definition wf_DP_R_xml_0_scc_31  := WF_DP_R_xml_0_scc_31.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_31 :
   forall x y, (DP_R_xml_0_scc_31 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_31).
    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_31.
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_32  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(U21(X1_,X2_)),U21(proper(X1_),proper(X2_))> *)
    | DP_R_xml_0_non_scc_32_0 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) x12) ->
       DP_R_xml_0_non_scc_32 (algebra.Alg.Term algebra.F.id_U21 
                              ((algebra.Alg.Term algebra.F.id_proper 
                              (x9::nil))::(algebra.Alg.Term 
                              algebra.F.id_proper (x10::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::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_31;
        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_33  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <U11(ok(X_)),U11(X_)> *)
    | DP_R_xml_0_scc_33_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
       DP_R_xml_0_scc_33 (algebra.Alg.Term algebra.F.id_U11 (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_U11 (x12::nil))
     (* <U11(mark(X_)),U11(X_)> *)
    | DP_R_xml_0_scc_33_1 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
       DP_R_xml_0_scc_33 (algebra.Alg.Term algebra.F.id_U11 (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_U11 (x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_33.
   Inductive DP_R_xml_0_scc_33_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U11(ok(X_)),U11(X_)> *)
     | DP_R_xml_0_scc_33_large_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
        DP_R_xml_0_scc_33_large (algebra.Alg.Term algebra.F.id_U11 (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_U11 (x12::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_33_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <U11(mark(X_)),U11(X_)> *)
     | DP_R_xml_0_scc_33_strict_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
        DP_R_xml_0_scc_33_strict (algebra.Alg.Term algebra.F.id_U11 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_U11 (x12::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_33_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 (x12:Z) := 2* x12.
    
    Definition P_id_U71 (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_isList (x12:Z) := 2 + 2* x12.
    
    Definition P_id_i  := 2.
    
    Definition P_id_U11 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_isQid (x12:Z) := 2 + 2* x12.
    
    Definition P_id_isNeList (x12:Z) := 2* x12.
    
    Definition P_id_ok (x12:Z) := 1 + 1* x12.
    
    Definition P_id_mark (x12:Z) := 0.
    
    Definition P_id_isPal (x12:Z) := 1* x12.
    
    Definition P_id_U41 (x12:Z) (x13:Z) := 1* x13.
    
    Definition P_id_u  := 1.
    
    Definition P_id_U21 (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_a  := 1.
    
    Definition P_id_U52 (x12:Z) := 2* x12.
    
    Definition P_id___ (x12:Z) (x13:Z) := 2* x13.
    
    Definition P_id_U72 (x12:Z) := 1* x12.
    
    Definition P_id_U31 (x12:Z) := 2 + 2* x12.
    
    Definition P_id_o  := 1.
    
    Definition P_id_tt  := 2.
    
    Definition P_id_isNePal (x12:Z) := 2* x12.
    
    Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
    
    Definition P_id_top (x12:Z) := 0.
    
    Definition P_id_nil  := 2.
    
    Definition P_id_U81 (x12:Z) := 2* x12.
    
    Definition P_id_U42 (x12:Z) := 1 + 2* x12.
    
    Definition P_id_proper (x12:Z) := 2* x12.
    
    Definition P_id_U22 (x12:Z) := 3 + 1* x12.
    
    Definition P_id_e  := 2.
    
    Definition P_id_U61 (x12:Z) := 1* x12.
    
    Lemma P_id_active_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_isList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
    Proof.
      intros x13 x12.
      intros [H_1 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 x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U72_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_top_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
    Proof.
      intros x13 x12.
      intros [H_1 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_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_active x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_i_bounded : 0 <= P_id_i .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
    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 x12, (0 <= x12) ->0 <= P_id_mark x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_u_bounded : 0 <= P_id_u .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_a_bounded : 0 <= P_id_a .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_o_bounded : 0 <= P_id_o .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_tt_bounded : 0 <= P_id_tt .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
    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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_proper x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_e_bounded : 0 <= P_id_e .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
    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_U71 P_id_isList P_id_i P_id_U11 
       P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
       P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
       P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
       P_id_proper P_id_U22 P_id_e P_id_U61.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                     P_id_active (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                     P_id_U71 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                     P_id_isList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                    | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                     P_id_U11 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                     P_id_isQid (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                     P_id_isNeList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                     P_id_ok (measure x12)
                    | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                     P_id_mark (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                     P_id_isPal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                     P_id_U41 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                    | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                     P_id_U21 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                    | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                     P_id_U52 (measure x12)
                    | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                     P_id___ (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                     P_id_U72 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                     P_id_U31 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                    | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                    | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                     P_id_isNePal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                     P_id_U51 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                     P_id_top (measure x12)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                     P_id_U81 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                     P_id_U42 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                     P_id_proper (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                     P_id_U22 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                    | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                     P_id_U61 (measure x12)
                    | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_U22_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNEPAL (x12:Z) := 0.
    
    Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U52_hat_1 (x12:Z) := 0.
    
    Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U42_hat_1 (x12:Z) := 0.
    
    Definition P_id_TOP (x12:Z) := 0.
    
    Definition P_id_ISQID (x12:Z) := 0.
    
    Definition P_id_ISPAL (x12:Z) := 0.
    
    Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_ACTIVE (x12:Z) := 0.
    
    Definition P_id_ISLIST (x12:Z) := 0.
    
    Definition P_id_PROPER (x12:Z) := 0.
    
    Definition P_id_U31_hat_1 (x12:Z) := 0.
    
    Definition P_id_U72_hat_1 (x12:Z) := 0.
    
    Definition P_id_U61_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNELIST (x12:Z) := 0.
    
    Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U11_hat_1 (x12:Z) := 1* x12.
    
    Definition P_id_U81_hat_1 (x12:Z) := 0.
    
    Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
    
    Lemma P_id_U22_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNEPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_TOP_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISQID_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISLIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_PROPER_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNELIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U11_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Definition marked_measure  := 
      InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
       P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
       P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
       P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
       P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
       P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
       P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
       P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
       P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
       P_id_U81_hat_1 P_id____hat_1.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                            P_id_U22_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNePal 
                              (x12::nil)) =>
                            P_id_ISNEPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U21 (x13::
                              x12::nil)) =>
                            P_id_U21_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                            P_id_U52_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U51 (x13::
                              x12::nil)) =>
                            P_id_U51_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                            P_id_U42_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                            P_id_TOP (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                            P_id_ISQID (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                            P_id_ISPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U71 (x13::
                              x12::nil)) =>
                            P_id_U71_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_active 
                              (x12::nil)) =>
                            P_id_ACTIVE (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isList 
                              (x12::nil)) =>
                            P_id_ISLIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_proper 
                              (x12::nil)) =>
                            P_id_PROPER (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                            P_id_U31_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                            P_id_U72_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                            P_id_U61_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNeList 
                              (x12::nil)) =>
                            P_id_ISNELIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U41 (x13::
                              x12::nil)) =>
                            P_id_U41_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                            P_id_U11_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                            P_id_U81_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id___ (x13::
                              x12::nil)) =>
                            P_id____hat_1 (measure x13) (measure x12)
                           | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_U22_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNEPAL_monotonic;assumption.
      intros ;apply P_id_U21_hat_1_monotonic;assumption.
      intros ;apply P_id_U52_hat_1_monotonic;assumption.
      intros ;apply P_id_U51_hat_1_monotonic;assumption.
      intros ;apply P_id_U42_hat_1_monotonic;assumption.
      intros ;apply P_id_TOP_monotonic;assumption.
      intros ;apply P_id_ISQID_monotonic;assumption.
      intros ;apply P_id_ISPAL_monotonic;assumption.
      intros ;apply P_id_U71_hat_1_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_ISLIST_monotonic;assumption.
      intros ;apply P_id_PROPER_monotonic;assumption.
      intros ;apply P_id_U31_hat_1_monotonic;assumption.
      intros ;apply P_id_U72_hat_1_monotonic;assumption.
      intros ;apply P_id_U61_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNELIST_monotonic;assumption.
      intros ;apply P_id_U41_hat_1_monotonic;assumption.
      intros ;apply P_id_U11_hat_1_monotonic;assumption.
      intros ;apply P_id_U81_hat_1_monotonic;assumption.
      intros ;apply P_id____hat_1_monotonic;assumption.
    Qed.
    
    Ltac rewrite_and_unfold  :=
     do 2 (rewrite marked_measure_equation);
      repeat (
      match goal with
        |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
         rewrite (measure_equation (algebra.Alg.Term f t))
        | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
         rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
        end
      ).
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_33.DP_R_xml_0_scc_33_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_33_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 (x12:Z) := 1 + 2* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 2* x12.
   
   Definition P_id_isList (x12:Z) := 0.
   
   Definition P_id_i  := 0.
   
   Definition P_id_U11 (x12:Z) := 1* x12.
   
   Definition P_id_isQid (x12:Z) := 0.
   
   Definition P_id_isNeList (x12:Z) := 0.
   
   Definition P_id_ok (x12:Z) := 1* x12.
   
   Definition P_id_mark (x12:Z) := 1 + 1* x12.
   
   Definition P_id_isPal (x12:Z) := 1.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_u  := 0.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_a  := 0.
   
   Definition P_id_U52 (x12:Z) := 1* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 2 + 1* x12 + 2* x13.
   
   Definition P_id_U72 (x12:Z) := 1 + 1* x12.
   
   Definition P_id_U31 (x12:Z) := 1* x12.
   
   Definition P_id_o  := 2.
   
   Definition P_id_tt  := 0.
   
   Definition P_id_isNePal (x12:Z) := 1.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_U81 (x12:Z) := 1* x12.
   
   Definition P_id_U42 (x12:Z) := 1* x12.
   
   Definition P_id_proper (x12:Z) := 2* x12.
   
   Definition P_id_U22 (x12:Z) := 1* x12.
   
   Definition P_id_e  := 0.
   
   Definition P_id_U61 (x12:Z) := 2 + 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 2* x12.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
   
   Definition le a b := marked_measure a <= marked_measure b.
   
   Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
   Proof.
     unfold lt, le in *.
     intros a b c.
     apply (interp.le_lt_compat_right (interp.o_Z 0)).
   Qed.
   
   Lemma wf_lt : well_founded lt.
   Proof.
     unfold lt in *.
     apply Inverse_Image.wf_inverse_image with  (B:=Z).
     apply Zwf.Zwf_well_founded.
   Qed.
   
   Lemma DP_R_xml_0_scc_33_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_33_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_33_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_33_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_33_large  := WF_DP_R_xml_0_scc_33_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_33.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_33_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_33_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_33_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_33_large.
   Qed.
  End WF_DP_R_xml_0_scc_33.
  
  Definition wf_DP_R_xml_0_scc_33  := WF_DP_R_xml_0_scc_33.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_33 :
   forall x y, (DP_R_xml_0_scc_33 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_33).
    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_33.
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_34  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(U11(X_)),U11(proper(X_))> *)
    | DP_R_xml_0_non_scc_34_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U11 (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_34 (algebra.Alg.Term algebra.F.id_U11 
                              ((algebra.Alg.Term algebra.F.id_proper 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::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_scc_33;
        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_35  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <__(X1_,mark(X2_)),__(X1_,X2_)> *)
    | DP_R_xml_0_scc_35_0 :
     forall x12 x10 x9 x13, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 x9 x13) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x10::nil)) x12) ->
        DP_R_xml_0_scc_35 (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) 
         (algebra.Alg.Term algebra.F.id___ (x13::x12::nil))
     (* <__(mark(X1_),X2_),__(X1_,X2_)> *)
    | DP_R_xml_0_scc_35_1 :
     forall x12 x10 x9 x13, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x9::nil)) x13) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x10 x12) ->
        DP_R_xml_0_scc_35 (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) 
         (algebra.Alg.Term algebra.F.id___ (x13::x12::nil))
     (* <__(ok(X1_),ok(X2_)),__(X1_,X2_)> *)
    | DP_R_xml_0_scc_35_2 :
     forall x12 x10 x9 x13, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x9::nil)) x13) ->
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x10::nil)) x12) ->
        DP_R_xml_0_scc_35 (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) 
         (algebra.Alg.Term algebra.F.id___ (x13::x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_35.
   Inductive DP_R_xml_0_scc_35_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <__(X1_,mark(X2_)),__(X1_,X2_)> *)
     | DP_R_xml_0_scc_35_large_0 :
      forall x12 x10 x9 x13, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  x9 x13) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_mark (x10::nil)) x12) ->
         DP_R_xml_0_scc_35_large (algebra.Alg.Term algebra.F.id___ (x9::
                                  x10::nil)) 
          (algebra.Alg.Term algebra.F.id___ (x13::x12::nil))
      (* <__(ok(X1_),ok(X2_)),__(X1_,X2_)> *)
     | DP_R_xml_0_scc_35_large_1 :
      forall x12 x10 x9 x13, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x9::nil)) x13) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_ok (x10::nil)) x12) ->
         DP_R_xml_0_scc_35_large (algebra.Alg.Term algebra.F.id___ (x9::
                                  x10::nil)) 
          (algebra.Alg.Term algebra.F.id___ (x13::x12::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_35_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <__(mark(X1_),X2_),__(X1_,X2_)> *)
     | DP_R_xml_0_scc_35_strict_0 :
      forall x12 x10 x9 x13, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x9::nil)) x13) ->
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   x10 x12) ->
         DP_R_xml_0_scc_35_strict (algebra.Alg.Term algebra.F.id___ (x9::
                                   x10::nil)) 
          (algebra.Alg.Term algebra.F.id___ (x13::x12::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_35_large.
    Inductive DP_R_xml_0_scc_35_large_large  :
     algebra.Alg.term ->algebra.Alg.term ->Prop := 
       (* <__(X1_,mark(X2_)),__(X1_,X2_)> *)
      | DP_R_xml_0_scc_35_large_large_0 :
       forall x12 x10 x9 x13, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   x9 x13) ->
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_mark (x10::nil)) x12) ->
          DP_R_xml_0_scc_35_large_large (algebra.Alg.Term algebra.F.id___ 
                                         (x9::x10::nil)) 
           (algebra.Alg.Term algebra.F.id___ (x13::x12::nil))
    .
    
    
    Inductive DP_R_xml_0_scc_35_large_strict  :
     algebra.Alg.term ->algebra.Alg.term ->Prop := 
       (* <__(ok(X1_),ok(X2_)),__(X1_,X2_)> *)
      | DP_R_xml_0_scc_35_large_strict_0 :
       forall x12 x10 x9 x13, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_ok (x9::nil)) x13) ->
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_ok (x10::nil)) x12) ->
          DP_R_xml_0_scc_35_large_strict (algebra.Alg.Term algebra.F.id___ 
                                          (x9::x10::nil)) 
           (algebra.Alg.Term algebra.F.id___ (x13::x12::nil))
    .
    
    
    Module WF_DP_R_xml_0_scc_35_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 (x12:Z) := 2* x12.
     
     Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 3* x12 + 1* x13.
     
     Definition P_id_isList (x12:Z) := 2 + 2* x12.
     
     Definition P_id_i  := 2.
     
     Definition P_id_U11 (x12:Z) := 1 + 1* x12.
     
     Definition P_id_isQid (x12:Z) := 2* x12.
     
     Definition P_id_isNeList (x12:Z) := 2 + 1* x12.
     
     Definition P_id_ok (x12:Z) := 0.
     
     Definition P_id_mark (x12:Z) := 1 + 1* x12.
     
     Definition P_id_isPal (x12:Z) := 3 + 2* x12.
     
     Definition P_id_U41 (x12:Z) (x13:Z) := 1 + 2* x12 + 2* x13.
     
     Definition P_id_u  := 2.
     
     Definition P_id_U21 (x12:Z) (x13:Z) := 2* x12 + 3* x13.
     
     Definition P_id_a  := 2.
     
     Definition P_id_U52 (x12:Z) := 1 + 3* x12.
     
     Definition P_id___ (x12:Z) (x13:Z) := 1 + 3* x12 + 2* x13.
     
     Definition P_id_U72 (x12:Z) := 3 + 1* x12.
     
     Definition P_id_U31 (x12:Z) := 1* x12.
     
     Definition P_id_o  := 2.
     
     Definition P_id_tt  := 2.
     
     Definition P_id_isNePal (x12:Z) := 2 + 2* x12.
     
     Definition P_id_U51 (x12:Z) (x13:Z) := 2* x12 + 3* x13.
     
     Definition P_id_top (x12:Z) := 0.
     
     Definition P_id_nil  := 0.
     
     Definition P_id_U81 (x12:Z) := 2* x12.
     
     Definition P_id_U42 (x12:Z) := 1 + 3* x12.
     
     Definition P_id_proper (x12:Z) := 1* x12.
     
     Definition P_id_U22 (x12:Z) := 1 + 2* x12.
     
     Definition P_id_e  := 2.
     
     Definition P_id_U61 (x12:Z) := 2* x12.
     
     Lemma P_id_active_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U71_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       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_isList_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U11_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isQid_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isNeList_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_ok_monotonic :
      forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
     Proof.
       intros x13 x12.
       intros [H_1 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 x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isPal_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U41_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       intros [H_1 H_0].
       intros [H_3 H_2].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U21_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       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_U52_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id____monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       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_U72_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U31_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isNePal_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U51_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       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_top_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U81_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U42_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_proper_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U22_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U61_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
     Proof.
       intros x13 x12.
       intros [H_1 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_bounded :
      forall x12, (0 <= x12) ->0 <= P_id_active x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U71_bounded :
      forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isList_bounded :
      forall x12, (0 <= x12) ->0 <= P_id_isList x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_i_bounded : 0 <= P_id_i .
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isNeList_bounded :
      forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
     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 x12, (0 <= x12) ->0 <= P_id_mark x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U41_bounded :
      forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_u_bounded : 0 <= P_id_u .
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U21_bounded :
      forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_a_bounded : 0 <= P_id_a .
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id____bounded :
      forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_o_bounded : 0 <= P_id_o .
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_tt_bounded : 0 <= P_id_tt .
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isNePal_bounded :
      forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U51_bounded :
      forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
     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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_proper_bounded :
      forall x12, (0 <= x12) ->0 <= P_id_proper x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_e_bounded : 0 <= P_id_e .
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
     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_U71 P_id_isList P_id_i P_id_U11 
        P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
        P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
        P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
        P_id_proper P_id_U22 P_id_e P_id_U61.
     
     Lemma measure_equation :
      forall t, 
       measure t = match t with
                     | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                      P_id_active (measure x12)
                     | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                      P_id_U71 (measure x13) (measure x12)
                     | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                      P_id_isList (measure x12)
                     | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                     | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                      P_id_U11 (measure x12)
                     | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                      P_id_isQid (measure x12)
                     | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                      P_id_isNeList (measure x12)
                     | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                      P_id_ok (measure x12)
                     | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                      P_id_mark (measure x12)
                     | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                      P_id_isPal (measure x12)
                     | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                      P_id_U41 (measure x13) (measure x12)
                     | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                     | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                      P_id_U21 (measure x13) (measure x12)
                     | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                     | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                      P_id_U52 (measure x12)
                     | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                      P_id___ (measure x13) (measure x12)
                     | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                      P_id_U72 (measure x12)
                     | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                      P_id_U31 (measure x12)
                     | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                     | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                     | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                      P_id_isNePal (measure x12)
                     | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                      P_id_U51 (measure x13) (measure x12)
                     | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                      P_id_top (measure x12)
                     | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                     | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                      P_id_U81 (measure x12)
                     | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                      P_id_U42 (measure x12)
                     | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                      P_id_proper (measure x12)
                     | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                      P_id_U22 (measure x12)
                     | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                     | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                      P_id_U61 (measure x12)
                     | _ => 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_U71_monotonic;assumption.
       intros ;apply P_id_isList_monotonic;assumption.
       intros ;apply P_id_U11_monotonic;assumption.
       intros ;apply P_id_isQid_monotonic;assumption.
       intros ;apply P_id_isNeList_monotonic;assumption.
       intros ;apply P_id_ok_monotonic;assumption.
       intros ;apply P_id_mark_monotonic;assumption.
       intros ;apply P_id_isPal_monotonic;assumption.
       intros ;apply P_id_U41_monotonic;assumption.
       intros ;apply P_id_U21_monotonic;assumption.
       intros ;apply P_id_U52_monotonic;assumption.
       intros ;apply P_id____monotonic;assumption.
       intros ;apply P_id_U72_monotonic;assumption.
       intros ;apply P_id_U31_monotonic;assumption.
       intros ;apply P_id_isNePal_monotonic;assumption.
       intros ;apply P_id_U51_monotonic;assumption.
       intros ;apply P_id_top_monotonic;assumption.
       intros ;apply P_id_U81_monotonic;assumption.
       intros ;apply P_id_U42_monotonic;assumption.
       intros ;apply P_id_proper_monotonic;assumption.
       intros ;apply P_id_U22_monotonic;assumption.
       intros ;apply P_id_U61_monotonic;assumption.
       intros ;apply P_id_active_bounded;assumption.
       intros ;apply P_id_U71_bounded;assumption.
       intros ;apply P_id_isList_bounded;assumption.
       intros ;apply P_id_i_bounded;assumption.
       intros ;apply P_id_U11_bounded;assumption.
       intros ;apply P_id_isQid_bounded;assumption.
       intros ;apply P_id_isNeList_bounded;assumption.
       intros ;apply P_id_ok_bounded;assumption.
       intros ;apply P_id_mark_bounded;assumption.
       intros ;apply P_id_isPal_bounded;assumption.
       intros ;apply P_id_U41_bounded;assumption.
       intros ;apply P_id_u_bounded;assumption.
       intros ;apply P_id_U21_bounded;assumption.
       intros ;apply P_id_a_bounded;assumption.
       intros ;apply P_id_U52_bounded;assumption.
       intros ;apply P_id____bounded;assumption.
       intros ;apply P_id_U72_bounded;assumption.
       intros ;apply P_id_U31_bounded;assumption.
       intros ;apply P_id_o_bounded;assumption.
       intros ;apply P_id_tt_bounded;assumption.
       intros ;apply P_id_isNePal_bounded;assumption.
       intros ;apply P_id_U51_bounded;assumption.
       intros ;apply P_id_top_bounded;assumption.
       intros ;apply P_id_nil_bounded;assumption.
       intros ;apply P_id_U81_bounded;assumption.
       intros ;apply P_id_U42_bounded;assumption.
       intros ;apply P_id_proper_bounded;assumption.
       intros ;apply P_id_U22_bounded;assumption.
       intros ;apply P_id_e_bounded;assumption.
       intros ;apply P_id_U61_bounded;assumption.
       apply rules_monotonic.
     Qed.
     
     Definition P_id_U22_hat_1 (x12:Z) := 0.
     
     Definition P_id_ISNEPAL (x12:Z) := 0.
     
     Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
     
     Definition P_id_U52_hat_1 (x12:Z) := 0.
     
     Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
     
     Definition P_id_U42_hat_1 (x12:Z) := 0.
     
     Definition P_id_TOP (x12:Z) := 0.
     
     Definition P_id_ISQID (x12:Z) := 0.
     
     Definition P_id_ISPAL (x12:Z) := 0.
     
     Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
     
     Definition P_id_ACTIVE (x12:Z) := 0.
     
     Definition P_id_ISLIST (x12:Z) := 0.
     
     Definition P_id_PROPER (x12:Z) := 0.
     
     Definition P_id_U31_hat_1 (x12:Z) := 0.
     
     Definition P_id_U72_hat_1 (x12:Z) := 0.
     
     Definition P_id_U61_hat_1 (x12:Z) := 0.
     
     Definition P_id_ISNELIST (x12:Z) := 0.
     
     Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
     
     Definition P_id_U11_hat_1 (x12:Z) := 0.
     
     Definition P_id_U81_hat_1 (x12:Z) := 0.
     
     Definition P_id____hat_1 (x12:Z) (x13:Z) := 1* x13.
     
     Lemma P_id_U22_hat_1_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_ISNEPAL_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U21_hat_1_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->
         P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       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_U52_hat_1_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U51_hat_1_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->
         P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       intros [H_1 H_0].
       intros [H_3 H_2].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U42_hat_1_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_TOP_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_ISQID_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_ISPAL_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U71_hat_1_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->
         P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       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_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_ISLIST_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_PROPER_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U31_hat_1_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U72_hat_1_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U61_hat_1_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_ISNELIST_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U41_hat_1_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->
         P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       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_U11_hat_1_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U81_hat_1_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id____hat_1_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->
         P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       intros [H_1 H_0].
       intros [H_3 H_2].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Definition marked_measure  := 
       InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
        P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
        P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
        P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
        P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
        P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
        P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
        P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
        P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
        P_id_U81_hat_1 P_id____hat_1.
     
     Lemma marked_measure_equation :
      forall t, 
       marked_measure t = match t with
                            | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                             P_id_U22_hat_1 (measure x12)
                            | (algebra.Alg.Term algebra.F.id_isNePal 
                               (x12::nil)) =>
                             P_id_ISNEPAL (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U21 (x13::
                               x12::nil)) =>
                             P_id_U21_hat_1 (measure x13) (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                             P_id_U52_hat_1 (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U51 (x13::
                               x12::nil)) =>
                             P_id_U51_hat_1 (measure x13) (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                             P_id_U42_hat_1 (measure x12)
                            | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                             P_id_TOP (measure x12)
                            | (algebra.Alg.Term algebra.F.id_isQid 
                               (x12::nil)) =>
                             P_id_ISQID (measure x12)
                            | (algebra.Alg.Term algebra.F.id_isPal 
                               (x12::nil)) =>
                             P_id_ISPAL (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U71 (x13::
                               x12::nil)) =>
                             P_id_U71_hat_1 (measure x13) (measure x12)
                            | (algebra.Alg.Term algebra.F.id_active 
                               (x12::nil)) =>
                             P_id_ACTIVE (measure x12)
                            | (algebra.Alg.Term algebra.F.id_isList 
                               (x12::nil)) =>
                             P_id_ISLIST (measure x12)
                            | (algebra.Alg.Term algebra.F.id_proper 
                               (x12::nil)) =>
                             P_id_PROPER (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                             P_id_U31_hat_1 (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                             P_id_U72_hat_1 (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                             P_id_U61_hat_1 (measure x12)
                            | (algebra.Alg.Term algebra.F.id_isNeList 
                               (x12::nil)) =>
                             P_id_ISNELIST (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U41 (x13::
                               x12::nil)) =>
                             P_id_U41_hat_1 (measure x13) (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                             P_id_U11_hat_1 (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                             P_id_U81_hat_1 (measure x12)
                            | (algebra.Alg.Term algebra.F.id___ (x13::
                               x12::nil)) =>
                             P_id____hat_1 (measure x13) (measure x12)
                            | _ => 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_U71_monotonic;assumption.
       intros ;apply P_id_isList_monotonic;assumption.
       intros ;apply P_id_U11_monotonic;assumption.
       intros ;apply P_id_isQid_monotonic;assumption.
       intros ;apply P_id_isNeList_monotonic;assumption.
       intros ;apply P_id_ok_monotonic;assumption.
       intros ;apply P_id_mark_monotonic;assumption.
       intros ;apply P_id_isPal_monotonic;assumption.
       intros ;apply P_id_U41_monotonic;assumption.
       intros ;apply P_id_U21_monotonic;assumption.
       intros ;apply P_id_U52_monotonic;assumption.
       intros ;apply P_id____monotonic;assumption.
       intros ;apply P_id_U72_monotonic;assumption.
       intros ;apply P_id_U31_monotonic;assumption.
       intros ;apply P_id_isNePal_monotonic;assumption.
       intros ;apply P_id_U51_monotonic;assumption.
       intros ;apply P_id_top_monotonic;assumption.
       intros ;apply P_id_U81_monotonic;assumption.
       intros ;apply P_id_U42_monotonic;assumption.
       intros ;apply P_id_proper_monotonic;assumption.
       intros ;apply P_id_U22_monotonic;assumption.
       intros ;apply P_id_U61_monotonic;assumption.
       intros ;apply P_id_active_bounded;assumption.
       intros ;apply P_id_U71_bounded;assumption.
       intros ;apply P_id_isList_bounded;assumption.
       intros ;apply P_id_i_bounded;assumption.
       intros ;apply P_id_U11_bounded;assumption.
       intros ;apply P_id_isQid_bounded;assumption.
       intros ;apply P_id_isNeList_bounded;assumption.
       intros ;apply P_id_ok_bounded;assumption.
       intros ;apply P_id_mark_bounded;assumption.
       intros ;apply P_id_isPal_bounded;assumption.
       intros ;apply P_id_U41_bounded;assumption.
       intros ;apply P_id_u_bounded;assumption.
       intros ;apply P_id_U21_bounded;assumption.
       intros ;apply P_id_a_bounded;assumption.
       intros ;apply P_id_U52_bounded;assumption.
       intros ;apply P_id____bounded;assumption.
       intros ;apply P_id_U72_bounded;assumption.
       intros ;apply P_id_U31_bounded;assumption.
       intros ;apply P_id_o_bounded;assumption.
       intros ;apply P_id_tt_bounded;assumption.
       intros ;apply P_id_isNePal_bounded;assumption.
       intros ;apply P_id_U51_bounded;assumption.
       intros ;apply P_id_top_bounded;assumption.
       intros ;apply P_id_nil_bounded;assumption.
       intros ;apply P_id_U81_bounded;assumption.
       intros ;apply P_id_U42_bounded;assumption.
       intros ;apply P_id_proper_bounded;assumption.
       intros ;apply P_id_U22_bounded;assumption.
       intros ;apply P_id_e_bounded;assumption.
       intros ;apply P_id_U61_bounded;assumption.
       apply rules_monotonic.
       intros ;apply P_id_U22_hat_1_monotonic;assumption.
       intros ;apply P_id_ISNEPAL_monotonic;assumption.
       intros ;apply P_id_U21_hat_1_monotonic;assumption.
       intros ;apply P_id_U52_hat_1_monotonic;assumption.
       intros ;apply P_id_U51_hat_1_monotonic;assumption.
       intros ;apply P_id_U42_hat_1_monotonic;assumption.
       intros ;apply P_id_TOP_monotonic;assumption.
       intros ;apply P_id_ISQID_monotonic;assumption.
       intros ;apply P_id_ISPAL_monotonic;assumption.
       intros ;apply P_id_U71_hat_1_monotonic;assumption.
       intros ;apply P_id_ACTIVE_monotonic;assumption.
       intros ;apply P_id_ISLIST_monotonic;assumption.
       intros ;apply P_id_PROPER_monotonic;assumption.
       intros ;apply P_id_U31_hat_1_monotonic;assumption.
       intros ;apply P_id_U72_hat_1_monotonic;assumption.
       intros ;apply P_id_U61_hat_1_monotonic;assumption.
       intros ;apply P_id_ISNELIST_monotonic;assumption.
       intros ;apply P_id_U41_hat_1_monotonic;assumption.
       intros ;apply P_id_U11_hat_1_monotonic;assumption.
       intros ;apply P_id_U81_hat_1_monotonic;assumption.
       intros ;apply P_id____hat_1_monotonic;assumption.
     Qed.
     
     Ltac rewrite_and_unfold  :=
      do 2 (rewrite marked_measure_equation);
       repeat (
       match goal with
         |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
          rewrite (measure_equation (algebra.Alg.Term f t))
         | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
          rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
         end
       ).
     
     
     Lemma wf :
      well_founded WF_DP_R_xml_0_scc_35_large.DP_R_xml_0_scc_35_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_35_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 (x12:Z) := 2* x12.
    
    Definition P_id_U71 (x12:Z) (x13:Z) := 3 + 1* x12 + 3* x13.
    
    Definition P_id_isList (x12:Z) := 2 + 2* x12.
    
    Definition P_id_i  := 2.
    
    Definition P_id_U11 (x12:Z) := 2 + 2* x12.
    
    Definition P_id_isQid (x12:Z) := 2* x12.
    
    Definition P_id_isNeList (x12:Z) := 2 + 2* x12.
    
    Definition P_id_ok (x12:Z) := 1 + 1* x12.
    
    Definition P_id_mark (x12:Z) := 0.
    
    Definition P_id_isPal (x12:Z) := 2 + 1* x12.
    
    Definition P_id_U41 (x12:Z) (x13:Z) := 2 + 3* x12 + 3* x13.
    
    Definition P_id_u  := 2.
    
    Definition P_id_U21 (x12:Z) (x13:Z) := 2 + 3* x12 + 3* x13.
    
    Definition P_id_a  := 2.
    
    Definition P_id_U52 (x12:Z) := 2 + 2* x12.
    
    Definition P_id___ (x12:Z) (x13:Z) := 2 + 3* x12 + 3* x13.
    
    Definition P_id_U72 (x12:Z) := 3 + 3* x12.
    
    Definition P_id_U31 (x12:Z) := 2 + 2* x12.
    
    Definition P_id_o  := 2.
    
    Definition P_id_tt  := 2.
    
    Definition P_id_isNePal (x12:Z) := 1* x12.
    
    Definition P_id_U51 (x12:Z) (x13:Z) := 2 + 3* x12 + 3* x13.
    
    Definition P_id_top (x12:Z) := 0.
    
    Definition P_id_nil  := 1.
    
    Definition P_id_U81 (x12:Z) := 2* x12.
    
    Definition P_id_U42 (x12:Z) := 2 + 2* x12.
    
    Definition P_id_proper (x12:Z) := 2* x12.
    
    Definition P_id_U22 (x12:Z) := 2 + 2* x12.
    
    Definition P_id_e  := 2.
    
    Definition P_id_U61 (x12:Z) := 2* x12.
    
    Lemma P_id_active_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_isList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
    Proof.
      intros x13 x12.
      intros [H_1 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 x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U72_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_top_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
    Proof.
      intros x13 x12.
      intros [H_1 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_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_active x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_i_bounded : 0 <= P_id_i .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
    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 x12, (0 <= x12) ->0 <= P_id_mark x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_u_bounded : 0 <= P_id_u .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_a_bounded : 0 <= P_id_a .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_o_bounded : 0 <= P_id_o .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_tt_bounded : 0 <= P_id_tt .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
    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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_proper x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_e_bounded : 0 <= P_id_e .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
    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_U71 P_id_isList P_id_i P_id_U11 
       P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
       P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
       P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
       P_id_proper P_id_U22 P_id_e P_id_U61.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                     P_id_active (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                     P_id_U71 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                     P_id_isList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                    | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                     P_id_U11 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                     P_id_isQid (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                     P_id_isNeList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                     P_id_ok (measure x12)
                    | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                     P_id_mark (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                     P_id_isPal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                     P_id_U41 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                    | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                     P_id_U21 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                    | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                     P_id_U52 (measure x12)
                    | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                     P_id___ (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                     P_id_U72 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                     P_id_U31 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                    | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                    | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                     P_id_isNePal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                     P_id_U51 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                     P_id_top (measure x12)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                     P_id_U81 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                     P_id_U42 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                     P_id_proper (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                     P_id_U22 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                    | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                     P_id_U61 (measure x12)
                    | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_U22_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNEPAL (x12:Z) := 0.
    
    Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U52_hat_1 (x12:Z) := 0.
    
    Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U42_hat_1 (x12:Z) := 0.
    
    Definition P_id_TOP (x12:Z) := 0.
    
    Definition P_id_ISQID (x12:Z) := 0.
    
    Definition P_id_ISPAL (x12:Z) := 0.
    
    Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_ACTIVE (x12:Z) := 0.
    
    Definition P_id_ISLIST (x12:Z) := 0.
    
    Definition P_id_PROPER (x12:Z) := 0.
    
    Definition P_id_U31_hat_1 (x12:Z) := 0.
    
    Definition P_id_U72_hat_1 (x12:Z) := 0.
    
    Definition P_id_U61_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNELIST (x12:Z) := 0.
    
    Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U11_hat_1 (x12:Z) := 0.
    
    Definition P_id_U81_hat_1 (x12:Z) := 0.
    
    Definition P_id____hat_1 (x12:Z) (x13:Z) := 1* x12.
    
    Lemma P_id_U22_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNEPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_TOP_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISQID_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISLIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_PROPER_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNELIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U11_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Definition marked_measure  := 
      InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
       P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
       P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
       P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
       P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
       P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
       P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
       P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
       P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
       P_id_U81_hat_1 P_id____hat_1.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                            P_id_U22_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNePal 
                              (x12::nil)) =>
                            P_id_ISNEPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U21 (x13::
                              x12::nil)) =>
                            P_id_U21_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                            P_id_U52_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U51 (x13::
                              x12::nil)) =>
                            P_id_U51_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                            P_id_U42_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                            P_id_TOP (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                            P_id_ISQID (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                            P_id_ISPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U71 (x13::
                              x12::nil)) =>
                            P_id_U71_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_active 
                              (x12::nil)) =>
                            P_id_ACTIVE (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isList 
                              (x12::nil)) =>
                            P_id_ISLIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_proper 
                              (x12::nil)) =>
                            P_id_PROPER (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                            P_id_U31_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                            P_id_U72_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                            P_id_U61_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNeList 
                              (x12::nil)) =>
                            P_id_ISNELIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U41 (x13::
                              x12::nil)) =>
                            P_id_U41_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                            P_id_U11_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                            P_id_U81_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id___ (x13::
                              x12::nil)) =>
                            P_id____hat_1 (measure x13) (measure x12)
                           | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_U22_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNEPAL_monotonic;assumption.
      intros ;apply P_id_U21_hat_1_monotonic;assumption.
      intros ;apply P_id_U52_hat_1_monotonic;assumption.
      intros ;apply P_id_U51_hat_1_monotonic;assumption.
      intros ;apply P_id_U42_hat_1_monotonic;assumption.
      intros ;apply P_id_TOP_monotonic;assumption.
      intros ;apply P_id_ISQID_monotonic;assumption.
      intros ;apply P_id_ISPAL_monotonic;assumption.
      intros ;apply P_id_U71_hat_1_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_ISLIST_monotonic;assumption.
      intros ;apply P_id_PROPER_monotonic;assumption.
      intros ;apply P_id_U31_hat_1_monotonic;assumption.
      intros ;apply P_id_U72_hat_1_monotonic;assumption.
      intros ;apply P_id_U61_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNELIST_monotonic;assumption.
      intros ;apply P_id_U41_hat_1_monotonic;assumption.
      intros ;apply P_id_U11_hat_1_monotonic;assumption.
      intros ;apply P_id_U81_hat_1_monotonic;assumption.
      intros ;apply P_id____hat_1_monotonic;assumption.
    Qed.
    
    Ltac rewrite_and_unfold  :=
     do 2 (rewrite marked_measure_equation);
      repeat (
      match goal with
        |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
         rewrite (measure_equation (algebra.Alg.Term f t))
        | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
         rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
        end
      ).
    
    Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
    
    Definition le a b := marked_measure a <= marked_measure b.
    
    Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
    Proof.
      unfold lt, le in *.
      intros a b c.
      apply (interp.le_lt_compat_right (interp.o_Z 0)).
    Qed.
    
    Lemma wf_lt : well_founded lt.
    Proof.
      unfold lt in *.
      apply Inverse_Image.wf_inverse_image with  (B:=Z).
      apply Zwf.Zwf_well_founded.
    Qed.
    
    Lemma DP_R_xml_0_scc_35_large_strict_in_lt :
     Relation_Definitions.inclusion _ DP_R_xml_0_scc_35_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_35_large_large_in_le :
     Relation_Definitions.inclusion _ DP_R_xml_0_scc_35_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_35_large_large  := 
      WF_DP_R_xml_0_scc_35_large_large.wf.
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_35.DP_R_xml_0_scc_35_large.
    Proof.
      intros x.
      apply (well_founded_ind wf_lt).
      clear x.
      intros x.
      pattern x.
      apply (@Acc_ind _ DP_R_xml_0_scc_35_large_large).
      clear x.
      intros x _ IHx IHx'.
      constructor.
      intros y H.
      
      destruct H;
       (apply IHx';apply DP_R_xml_0_scc_35_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_35_large_large_in_le;econstructor eassumption])).
      apply wf_DP_R_xml_0_scc_35_large_large.
    Qed.
   End WF_DP_R_xml_0_scc_35_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 (x12:Z) := 1 + 1* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_isList (x12:Z) := 0.
   
   Definition P_id_i  := 0.
   
   Definition P_id_U11 (x12:Z) := 1* x12.
   
   Definition P_id_isQid (x12:Z) := 0.
   
   Definition P_id_isNeList (x12:Z) := 0.
   
   Definition P_id_ok (x12:Z) := 2* x12.
   
   Definition P_id_mark (x12:Z) := 1 + 1* x12.
   
   Definition P_id_isPal (x12:Z) := 0.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_u  := 0.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_a  := 0.
   
   Definition P_id_U52 (x12:Z) := 1* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 1* x12 + 1* x13.
   
   Definition P_id_U72 (x12:Z) := 1* x12.
   
   Definition P_id_U31 (x12:Z) := 1* x12.
   
   Definition P_id_o  := 0.
   
   Definition P_id_tt  := 0.
   
   Definition P_id_isNePal (x12:Z) := 0.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 1* x12.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_U81 (x12:Z) := 1* x12.
   
   Definition P_id_U42 (x12:Z) := 1* x12.
   
   Definition P_id_proper (x12:Z) := 0.
   
   Definition P_id_U22 (x12:Z) := 1* x12.
   
   Definition P_id_e  := 0.
   
   Definition P_id_U61 (x12:Z) := 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 1* x12.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
   
   Definition le a b := marked_measure a <= marked_measure b.
   
   Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
   Proof.
     unfold lt, le in *.
     intros a b c.
     apply (interp.le_lt_compat_right (interp.o_Z 0)).
   Qed.
   
   Lemma wf_lt : well_founded lt.
   Proof.
     unfold lt in *.
     apply Inverse_Image.wf_inverse_image with  (B:=Z).
     apply Zwf.Zwf_well_founded.
   Qed.
   
   Lemma DP_R_xml_0_scc_35_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_35_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_35_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_35_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_35_large  := WF_DP_R_xml_0_scc_35_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_35.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_35_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_35_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_35_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_35_large.
   Qed.
  End WF_DP_R_xml_0_scc_35.
  
  Definition wf_DP_R_xml_0_scc_35  := WF_DP_R_xml_0_scc_35.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_35 :
   forall x y, (DP_R_xml_0_scc_35 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_35).
    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_35.
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_36  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(__(X1_,X2_)),__(proper(X1_),proper(X2_))> *)
    | DP_R_xml_0_non_scc_36_0 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
       DP_R_xml_0_non_scc_36 (algebra.Alg.Term algebra.F.id___ 
                              ((algebra.Alg.Term algebra.F.id_proper 
                              (x9::nil))::(algebra.Alg.Term 
                              algebra.F.id_proper (x10::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::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_scc_35;
        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_37  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <proper(__(X1_,X2_)),proper(X2_)> *)
    | DP_R_xml_0_scc_37_0 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x10::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(__(X1_,X2_)),proper(X1_)> *)
    | DP_R_xml_0_scc_37_1 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x9::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(U11(X_)),proper(X_)> *)
    | DP_R_xml_0_scc_37_2 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U11 (x1::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(U21(X1_,X2_)),proper(X1_)> *)
    | DP_R_xml_0_scc_37_3 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x9::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(U21(X1_,X2_)),proper(X2_)> *)
    | DP_R_xml_0_scc_37_4 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x10::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(U22(X_)),proper(X_)> *)
    | DP_R_xml_0_scc_37_5 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U22 (x1::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(isList(X_)),proper(X_)> *)
    | DP_R_xml_0_scc_37_6 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(U31(X_)),proper(X_)> *)
    | DP_R_xml_0_scc_37_7 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U31 (x1::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(U41(X1_,X2_)),proper(X1_)> *)
    | DP_R_xml_0_scc_37_8 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U41 (x9::x10::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x9::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(U41(X1_,X2_)),proper(X2_)> *)
    | DP_R_xml_0_scc_37_9 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U41 (x9::x10::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x10::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(U42(X_)),proper(X_)> *)
    | DP_R_xml_0_scc_37_10 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U42 (x1::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(isNeList(X_)),proper(X_)> *)
    | DP_R_xml_0_scc_37_11 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(U51(X1_,X2_)),proper(X1_)> *)
    | DP_R_xml_0_scc_37_12 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U51 (x9::x10::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x9::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(U51(X1_,X2_)),proper(X2_)> *)
    | DP_R_xml_0_scc_37_13 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U51 (x9::x10::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x10::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(U52(X_)),proper(X_)> *)
    | DP_R_xml_0_scc_37_14 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U52 (x1::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(U61(X_)),proper(X_)> *)
    | DP_R_xml_0_scc_37_15 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U61 (x1::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(U71(X1_,X2_)),proper(X1_)> *)
    | DP_R_xml_0_scc_37_16 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x9::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(U71(X1_,X2_)),proper(X2_)> *)
    | DP_R_xml_0_scc_37_17 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x10::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(U72(X_)),proper(X_)> *)
    | DP_R_xml_0_scc_37_18 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(isPal(X_)),proper(X_)> *)
    | DP_R_xml_0_scc_37_19 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isPal (x1::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(U81(X_)),proper(X_)> *)
    | DP_R_xml_0_scc_37_20 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U81 (x1::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(isQid(X_)),proper(X_)> *)
    | DP_R_xml_0_scc_37_21 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isQid (x1::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     (* <proper(isNePal(X_)),proper(X_)> *)
    | DP_R_xml_0_scc_37_22 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isNePal (x1::nil)) x12) ->
       DP_R_xml_0_scc_37 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_proper (x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_37.
   Inductive DP_R_xml_0_scc_37_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <proper(__(X1_,X2_)),proper(X2_)> *)
     | DP_R_xml_0_scc_37_large_0 :
      forall x12 x10 x9, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
        DP_R_xml_0_scc_37_large (algebra.Alg.Term algebra.F.id_proper 
                                 (x10::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(__(X1_,X2_)),proper(X1_)> *)
     | DP_R_xml_0_scc_37_large_1 :
      forall x12 x10 x9, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
        DP_R_xml_0_scc_37_large (algebra.Alg.Term algebra.F.id_proper 
                                 (x9::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(U11(X_)),proper(X_)> *)
     | DP_R_xml_0_scc_37_large_2 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_U11 (x1::nil)) x12) ->
        DP_R_xml_0_scc_37_large (algebra.Alg.Term algebra.F.id_proper 
                                 (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(U21(X1_,X2_)),proper(X1_)> *)
     | DP_R_xml_0_scc_37_large_3 :
      forall x12 x10 x9, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) x12) ->
        DP_R_xml_0_scc_37_large (algebra.Alg.Term algebra.F.id_proper 
                                 (x9::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(U21(X1_,X2_)),proper(X2_)> *)
     | DP_R_xml_0_scc_37_large_4 :
      forall x12 x10 x9, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) x12) ->
        DP_R_xml_0_scc_37_large (algebra.Alg.Term algebra.F.id_proper 
                                 (x10::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(isList(X_)),proper(X_)> *)
     | DP_R_xml_0_scc_37_large_5 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x12) ->
        DP_R_xml_0_scc_37_large (algebra.Alg.Term algebra.F.id_proper 
                                 (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_37_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <proper(U22(X_)),proper(X_)> *)
     | DP_R_xml_0_scc_37_strict_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_U22 (x1::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(U31(X_)),proper(X_)> *)
     | DP_R_xml_0_scc_37_strict_1 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_U31 (x1::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(U41(X1_,X2_)),proper(X1_)> *)
     | DP_R_xml_0_scc_37_strict_2 :
      forall x12 x10 x9, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_U41 (x9::x10::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x9::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(U41(X1_,X2_)),proper(X2_)> *)
     | DP_R_xml_0_scc_37_strict_3 :
      forall x12 x10 x9, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_U41 (x9::x10::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x10::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(U42(X_)),proper(X_)> *)
     | DP_R_xml_0_scc_37_strict_4 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_U42 (x1::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(isNeList(X_)),proper(X_)> *)
     | DP_R_xml_0_scc_37_strict_5 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(U51(X1_,X2_)),proper(X1_)> *)
     | DP_R_xml_0_scc_37_strict_6 :
      forall x12 x10 x9, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_U51 (x9::x10::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x9::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(U51(X1_,X2_)),proper(X2_)> *)
     | DP_R_xml_0_scc_37_strict_7 :
      forall x12 x10 x9, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_U51 (x9::x10::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x10::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(U52(X_)),proper(X_)> *)
     | DP_R_xml_0_scc_37_strict_8 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_U52 (x1::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(U61(X_)),proper(X_)> *)
     | DP_R_xml_0_scc_37_strict_9 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_U61 (x1::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(U71(X1_,X2_)),proper(X1_)> *)
     | DP_R_xml_0_scc_37_strict_10 :
      forall x12 x10 x9, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x9::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(U71(X1_,X2_)),proper(X2_)> *)
     | DP_R_xml_0_scc_37_strict_11 :
      forall x12 x10 x9, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x10::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(U72(X_)),proper(X_)> *)
     | DP_R_xml_0_scc_37_strict_12 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(isPal(X_)),proper(X_)> *)
     | DP_R_xml_0_scc_37_strict_13 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_isPal (x1::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(U81(X_)),proper(X_)> *)
     | DP_R_xml_0_scc_37_strict_14 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_U81 (x1::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(isQid(X_)),proper(X_)> *)
     | DP_R_xml_0_scc_37_strict_15 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_isQid (x1::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      (* <proper(isNePal(X_)),proper(X_)> *)
     | DP_R_xml_0_scc_37_strict_16 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_isNePal (x1::nil)) x12) ->
        DP_R_xml_0_scc_37_strict (algebra.Alg.Term algebra.F.id_proper 
                                  (x1::nil)) 
         (algebra.Alg.Term algebra.F.id_proper (x12::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_37_large.
    Inductive DP_R_xml_0_scc_37_large_large  :
     algebra.Alg.term ->algebra.Alg.term ->Prop := 
       (* <proper(__(X1_,X2_)),proper(X2_)> *)
      | DP_R_xml_0_scc_37_large_large_0 :
       forall x12 x10 x9, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
         DP_R_xml_0_scc_37_large_large (algebra.Alg.Term algebra.F.id_proper 
                                        (x10::nil)) 
          (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      
       (* <proper(__(X1_,X2_)),proper(X1_)> *)
      | DP_R_xml_0_scc_37_large_large_1 :
       forall x12 x10 x9, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
         DP_R_xml_0_scc_37_large_large (algebra.Alg.Term algebra.F.id_proper 
                                        (x9::nil)) 
          (algebra.Alg.Term algebra.F.id_proper (x12::nil))
       (* <proper(U11(X_)),proper(X_)> *)
      | DP_R_xml_0_scc_37_large_large_2 :
       forall x12 x1, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_U11 (x1::nil)) x12) ->
         DP_R_xml_0_scc_37_large_large (algebra.Alg.Term algebra.F.id_proper 
                                        (x1::nil)) 
          (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    .
    
    
    Inductive DP_R_xml_0_scc_37_large_strict  :
     algebra.Alg.term ->algebra.Alg.term ->Prop := 
       (* <proper(U21(X1_,X2_)),proper(X1_)> *)
      | DP_R_xml_0_scc_37_large_strict_0 :
       forall x12 x10 x9, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) x12) ->
         DP_R_xml_0_scc_37_large_strict (algebra.Alg.Term 
                                         algebra.F.id_proper (x9::nil)) 
          (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      
       (* <proper(U21(X1_,X2_)),proper(X2_)> *)
      | DP_R_xml_0_scc_37_large_strict_1 :
       forall x12 x10 x9, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) x12) ->
         DP_R_xml_0_scc_37_large_strict (algebra.Alg.Term 
                                         algebra.F.id_proper (x10::nil)) 
          (algebra.Alg.Term algebra.F.id_proper (x12::nil))
      
       (* <proper(isList(X_)),proper(X_)> *)
      | DP_R_xml_0_scc_37_large_strict_2 :
       forall x12 x1, 
        (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                   
          (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x12) ->
         DP_R_xml_0_scc_37_large_strict (algebra.Alg.Term 
                                         algebra.F.id_proper (x1::nil)) 
          (algebra.Alg.Term algebra.F.id_proper (x12::nil))
    .
    
    
    Module WF_DP_R_xml_0_scc_37_large_large.
     Inductive DP_R_xml_0_scc_37_large_large_large  :
      algebra.Alg.term ->algebra.Alg.term ->Prop := 
        (* <proper(U11(X_)),proper(X_)> *)
       | DP_R_xml_0_scc_37_large_large_large_0 :
        forall x12 x1, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id_U11 (x1::nil)) x12) ->
          DP_R_xml_0_scc_37_large_large_large (algebra.Alg.Term 
                                               algebra.F.id_proper (x1::nil)) 
           (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     .
     
     
     Inductive DP_R_xml_0_scc_37_large_large_strict  :
      algebra.Alg.term ->algebra.Alg.term ->Prop := 
        (* <proper(__(X1_,X2_)),proper(X2_)> *)
       | DP_R_xml_0_scc_37_large_large_strict_0 :
        forall x12 x10 x9, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
          DP_R_xml_0_scc_37_large_large_strict (algebra.Alg.Term 
                                                algebra.F.id_proper 
                                                (x10::nil)) 
           (algebra.Alg.Term algebra.F.id_proper (x12::nil))
       
        (* <proper(__(X1_,X2_)),proper(X1_)> *)
       | DP_R_xml_0_scc_37_large_large_strict_1 :
        forall x12 x10 x9, 
         (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                    
           (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
          DP_R_xml_0_scc_37_large_large_strict (algebra.Alg.Term 
                                                algebra.F.id_proper 
                                                (x9::nil)) 
           (algebra.Alg.Term algebra.F.id_proper (x12::nil))
     .
     
     
     Module WF_DP_R_xml_0_scc_37_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 (x12:Z) := 1 + 2* x12.
      
      Definition P_id_U71 (x12:Z) (x13:Z) := 0.
      
      Definition P_id_isList (x12:Z) := 3 + 2* x12.
      
      Definition P_id_i  := 3.
      
      Definition P_id_U11 (x12:Z) := 1 + 1* x12.
      
      Definition P_id_isQid (x12:Z) := 3* x12.
      
      Definition P_id_isNeList (x12:Z) := 3.
      
      Definition P_id_ok (x12:Z) := 0.
      
      Definition P_id_mark (x12:Z) := 0.
      
      Definition P_id_isPal (x12:Z) := 3* x12.
      
      Definition P_id_U41 (x12:Z) (x13:Z) := 3 + 1* x12 + 3* x13.
      
      Definition P_id_u  := 3.
      
      Definition P_id_U21 (x12:Z) (x13:Z) := 3 + 1* x12 + 3* x13.
      
      Definition P_id_a  := 3.
      
      Definition P_id_U52 (x12:Z) := 1 + 1* x12.
      
      Definition P_id___ (x12:Z) (x13:Z) := 1* x12.
      
      Definition P_id_U72 (x12:Z) := 0.
      
      Definition P_id_U31 (x12:Z) := 3 + 1* x12.
      
      Definition P_id_o  := 3.
      
      Definition P_id_tt  := 0.
      
      Definition P_id_isNePal (x12:Z) := 0.
      
      Definition P_id_U51 (x12:Z) (x13:Z) := 3* x13.
      
      Definition P_id_top (x12:Z) := 0.
      
      Definition P_id_nil  := 3.
      
      Definition P_id_U81 (x12:Z) := 0.
      
      Definition P_id_U42 (x12:Z) := 3 + 1* x12.
      
      Definition P_id_proper (x12:Z) := 1* x12.
      
      Definition P_id_U22 (x12:Z) := 0.
      
      Definition P_id_e  := 3.
      
      Definition P_id_U61 (x12:Z) := 1 + 1* x12.
      
      Lemma P_id_active_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U71_monotonic :
       forall x12 x14 x13 x15, 
        (0 <= x15)/\ (x15 <= x14) ->
         (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
      Proof.
        intros x15 x14 x13 x12.
        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_isList_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U11_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_isQid_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_isNeList_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_ok_monotonic :
       forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
      Proof.
        intros x13 x12.
        intros [H_1 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 x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_isPal_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U41_monotonic :
       forall x12 x14 x13 x15, 
        (0 <= x15)/\ (x15 <= x14) ->
         (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
      Proof.
        intros x15 x14 x13 x12.
        intros [H_1 H_0].
        intros [H_3 H_2].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U21_monotonic :
       forall x12 x14 x13 x15, 
        (0 <= x15)/\ (x15 <= x14) ->
         (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
      Proof.
        intros x15 x14 x13 x12.
        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_U52_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id____monotonic :
       forall x12 x14 x13 x15, 
        (0 <= x15)/\ (x15 <= x14) ->
         (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
      Proof.
        intros x15 x14 x13 x12.
        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_U72_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U31_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_isNePal_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U51_monotonic :
       forall x12 x14 x13 x15, 
        (0 <= x15)/\ (x15 <= x14) ->
         (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
      Proof.
        intros x15 x14 x13 x12.
        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_top_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U81_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U42_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_proper_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U22_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U61_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
      Proof.
        intros x13 x12.
        intros [H_1 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_bounded :
       forall x12, (0 <= x12) ->0 <= P_id_active x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U71_bounded :
       forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_isList_bounded :
       forall x12, (0 <= x12) ->0 <= P_id_isList x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_i_bounded : 0 <= P_id_i .
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_isQid_bounded :
       forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_isNeList_bounded :
       forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
      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 x12, (0 <= x12) ->0 <= P_id_mark x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_isPal_bounded :
       forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U41_bounded :
       forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_u_bounded : 0 <= P_id_u .
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U21_bounded :
       forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_a_bounded : 0 <= P_id_a .
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id____bounded :
       forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_o_bounded : 0 <= P_id_o .
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_tt_bounded : 0 <= P_id_tt .
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_isNePal_bounded :
       forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U51_bounded :
       forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
      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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_proper_bounded :
       forall x12, (0 <= x12) ->0 <= P_id_proper x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_e_bounded : 0 <= P_id_e .
      Proof.
        intros .
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
      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_U71 P_id_isList P_id_i P_id_U11 
         P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
         P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
         P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
         P_id_proper P_id_U22 P_id_e P_id_U61.
      
      Lemma measure_equation :
       forall t, 
        measure t = match t with
                      | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                       P_id_active (measure x12)
                      | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                       P_id_U71 (measure x13) (measure x12)
                      | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                       P_id_isList (measure x12)
                      | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                      | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                       P_id_U11 (measure x12)
                      | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                       P_id_isQid (measure x12)
                      | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                       P_id_isNeList (measure x12)
                      | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                       P_id_ok (measure x12)
                      | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                       P_id_mark (measure x12)
                      | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                       P_id_isPal (measure x12)
                      | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                       P_id_U41 (measure x13) (measure x12)
                      | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                      | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                       P_id_U21 (measure x13) (measure x12)
                      | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                      | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                       P_id_U52 (measure x12)
                      | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                       P_id___ (measure x13) (measure x12)
                      | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                       P_id_U72 (measure x12)
                      | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                       P_id_U31 (measure x12)
                      | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                      | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                      | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                       P_id_isNePal (measure x12)
                      | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                       P_id_U51 (measure x13) (measure x12)
                      | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                       P_id_top (measure x12)
                      | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                      | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                       P_id_U81 (measure x12)
                      | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                       P_id_U42 (measure x12)
                      | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                       P_id_proper (measure x12)
                      | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                       P_id_U22 (measure x12)
                      | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                      | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                       P_id_U61 (measure x12)
                      | _ => 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_U71_monotonic;assumption.
        intros ;apply P_id_isList_monotonic;assumption.
        intros ;apply P_id_U11_monotonic;assumption.
        intros ;apply P_id_isQid_monotonic;assumption.
        intros ;apply P_id_isNeList_monotonic;assumption.
        intros ;apply P_id_ok_monotonic;assumption.
        intros ;apply P_id_mark_monotonic;assumption.
        intros ;apply P_id_isPal_monotonic;assumption.
        intros ;apply P_id_U41_monotonic;assumption.
        intros ;apply P_id_U21_monotonic;assumption.
        intros ;apply P_id_U52_monotonic;assumption.
        intros ;apply P_id____monotonic;assumption.
        intros ;apply P_id_U72_monotonic;assumption.
        intros ;apply P_id_U31_monotonic;assumption.
        intros ;apply P_id_isNePal_monotonic;assumption.
        intros ;apply P_id_U51_monotonic;assumption.
        intros ;apply P_id_top_monotonic;assumption.
        intros ;apply P_id_U81_monotonic;assumption.
        intros ;apply P_id_U42_monotonic;assumption.
        intros ;apply P_id_proper_monotonic;assumption.
        intros ;apply P_id_U22_monotonic;assumption.
        intros ;apply P_id_U61_monotonic;assumption.
        intros ;apply P_id_active_bounded;assumption.
        intros ;apply P_id_U71_bounded;assumption.
        intros ;apply P_id_isList_bounded;assumption.
        intros ;apply P_id_i_bounded;assumption.
        intros ;apply P_id_U11_bounded;assumption.
        intros ;apply P_id_isQid_bounded;assumption.
        intros ;apply P_id_isNeList_bounded;assumption.
        intros ;apply P_id_ok_bounded;assumption.
        intros ;apply P_id_mark_bounded;assumption.
        intros ;apply P_id_isPal_bounded;assumption.
        intros ;apply P_id_U41_bounded;assumption.
        intros ;apply P_id_u_bounded;assumption.
        intros ;apply P_id_U21_bounded;assumption.
        intros ;apply P_id_a_bounded;assumption.
        intros ;apply P_id_U52_bounded;assumption.
        intros ;apply P_id____bounded;assumption.
        intros ;apply P_id_U72_bounded;assumption.
        intros ;apply P_id_U31_bounded;assumption.
        intros ;apply P_id_o_bounded;assumption.
        intros ;apply P_id_tt_bounded;assumption.
        intros ;apply P_id_isNePal_bounded;assumption.
        intros ;apply P_id_U51_bounded;assumption.
        intros ;apply P_id_top_bounded;assumption.
        intros ;apply P_id_nil_bounded;assumption.
        intros ;apply P_id_U81_bounded;assumption.
        intros ;apply P_id_U42_bounded;assumption.
        intros ;apply P_id_proper_bounded;assumption.
        intros ;apply P_id_U22_bounded;assumption.
        intros ;apply P_id_e_bounded;assumption.
        intros ;apply P_id_U61_bounded;assumption.
        apply rules_monotonic.
      Qed.
      
      Definition P_id_U22_hat_1 (x12:Z) := 0.
      
      Definition P_id_ISNEPAL (x12:Z) := 0.
      
      Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
      
      Definition P_id_U52_hat_1 (x12:Z) := 0.
      
      Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
      
      Definition P_id_U42_hat_1 (x12:Z) := 0.
      
      Definition P_id_TOP (x12:Z) := 0.
      
      Definition P_id_ISQID (x12:Z) := 0.
      
      Definition P_id_ISPAL (x12:Z) := 0.
      
      Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
      
      Definition P_id_ACTIVE (x12:Z) := 0.
      
      Definition P_id_ISLIST (x12:Z) := 0.
      
      Definition P_id_PROPER (x12:Z) := 1* x12.
      
      Definition P_id_U31_hat_1 (x12:Z) := 0.
      
      Definition P_id_U72_hat_1 (x12:Z) := 0.
      
      Definition P_id_U61_hat_1 (x12:Z) := 0.
      
      Definition P_id_ISNELIST (x12:Z) := 0.
      
      Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
      
      Definition P_id_U11_hat_1 (x12:Z) := 0.
      
      Definition P_id_U81_hat_1 (x12:Z) := 0.
      
      Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
      
      Lemma P_id_U22_hat_1_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_ISNEPAL_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U21_hat_1_monotonic :
       forall x12 x14 x13 x15, 
        (0 <= x15)/\ (x15 <= x14) ->
         (0 <= x13)/\ (x13 <= x12) ->
          P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
      Proof.
        intros x15 x14 x13 x12.
        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_U52_hat_1_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U51_hat_1_monotonic :
       forall x12 x14 x13 x15, 
        (0 <= x15)/\ (x15 <= x14) ->
         (0 <= x13)/\ (x13 <= x12) ->
          P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
      Proof.
        intros x15 x14 x13 x12.
        intros [H_1 H_0].
        intros [H_3 H_2].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U42_hat_1_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_TOP_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_ISQID_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_ISPAL_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U71_hat_1_monotonic :
       forall x12 x14 x13 x15, 
        (0 <= x15)/\ (x15 <= x14) ->
         (0 <= x13)/\ (x13 <= x12) ->
          P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
      Proof.
        intros x15 x14 x13 x12.
        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_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_ISLIST_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_PROPER_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U31_hat_1_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U72_hat_1_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U61_hat_1_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_ISNELIST_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U41_hat_1_monotonic :
       forall x12 x14 x13 x15, 
        (0 <= x15)/\ (x15 <= x14) ->
         (0 <= x13)/\ (x13 <= x12) ->
          P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
      Proof.
        intros x15 x14 x13 x12.
        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_U11_hat_1_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id_U81_hat_1_monotonic :
       forall x12 x13, 
        (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
      Proof.
        intros x13 x12.
        intros [H_1 H_0].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Lemma P_id____hat_1_monotonic :
       forall x12 x14 x13 x15, 
        (0 <= x15)/\ (x15 <= x14) ->
         (0 <= x13)/\ (x13 <= x12) ->
          P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
      Proof.
        intros x15 x14 x13 x12.
        intros [H_1 H_0].
        intros [H_3 H_2].
        
        cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
         (auto with zarith)||(repeat (translate_vars );prove_ineq ).
      Qed.
      
      Definition marked_measure  := 
        InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
         P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
         P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
         P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
         P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
         P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
         P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
         P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
         P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
         P_id_U81_hat_1 P_id____hat_1.
      
      Lemma marked_measure_equation :
       forall t, 
        marked_measure t = match t with
                             | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                              P_id_U22_hat_1 (measure x12)
                             | (algebra.Alg.Term algebra.F.id_isNePal 
                                (x12::nil)) =>
                              P_id_ISNEPAL (measure x12)
                             | (algebra.Alg.Term algebra.F.id_U21 (x13::
                                x12::nil)) =>
                              P_id_U21_hat_1 (measure x13) (measure x12)
                             | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                              P_id_U52_hat_1 (measure x12)
                             | (algebra.Alg.Term algebra.F.id_U51 (x13::
                                x12::nil)) =>
                              P_id_U51_hat_1 (measure x13) (measure x12)
                             | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                              P_id_U42_hat_1 (measure x12)
                             | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                              P_id_TOP (measure x12)
                             | (algebra.Alg.Term algebra.F.id_isQid 
                                (x12::nil)) =>
                              P_id_ISQID (measure x12)
                             | (algebra.Alg.Term algebra.F.id_isPal 
                                (x12::nil)) =>
                              P_id_ISPAL (measure x12)
                             | (algebra.Alg.Term algebra.F.id_U71 (x13::
                                x12::nil)) =>
                              P_id_U71_hat_1 (measure x13) (measure x12)
                             | (algebra.Alg.Term algebra.F.id_active 
                                (x12::nil)) =>
                              P_id_ACTIVE (measure x12)
                             | (algebra.Alg.Term algebra.F.id_isList 
                                (x12::nil)) =>
                              P_id_ISLIST (measure x12)
                             | (algebra.Alg.Term algebra.F.id_proper 
                                (x12::nil)) =>
                              P_id_PROPER (measure x12)
                             | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                              P_id_U31_hat_1 (measure x12)
                             | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                              P_id_U72_hat_1 (measure x12)
                             | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                              P_id_U61_hat_1 (measure x12)
                             | (algebra.Alg.Term algebra.F.id_isNeList 
                                (x12::nil)) =>
                              P_id_ISNELIST (measure x12)
                             | (algebra.Alg.Term algebra.F.id_U41 (x13::
                                x12::nil)) =>
                              P_id_U41_hat_1 (measure x13) (measure x12)
                             | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                              P_id_U11_hat_1 (measure x12)
                             | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                              P_id_U81_hat_1 (measure x12)
                             | (algebra.Alg.Term algebra.F.id___ (x13::
                                x12::nil)) =>
                              P_id____hat_1 (measure x13) (measure x12)
                             | _ => 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_U71_monotonic;assumption.
        intros ;apply P_id_isList_monotonic;assumption.
        intros ;apply P_id_U11_monotonic;assumption.
        intros ;apply P_id_isQid_monotonic;assumption.
        intros ;apply P_id_isNeList_monotonic;assumption.
        intros ;apply P_id_ok_monotonic;assumption.
        intros ;apply P_id_mark_monotonic;assumption.
        intros ;apply P_id_isPal_monotonic;assumption.
        intros ;apply P_id_U41_monotonic;assumption.
        intros ;apply P_id_U21_monotonic;assumption.
        intros ;apply P_id_U52_monotonic;assumption.
        intros ;apply P_id____monotonic;assumption.
        intros ;apply P_id_U72_monotonic;assumption.
        intros ;apply P_id_U31_monotonic;assumption.
        intros ;apply P_id_isNePal_monotonic;assumption.
        intros ;apply P_id_U51_monotonic;assumption.
        intros ;apply P_id_top_monotonic;assumption.
        intros ;apply P_id_U81_monotonic;assumption.
        intros ;apply P_id_U42_monotonic;assumption.
        intros ;apply P_id_proper_monotonic;assumption.
        intros ;apply P_id_U22_monotonic;assumption.
        intros ;apply P_id_U61_monotonic;assumption.
        intros ;apply P_id_active_bounded;assumption.
        intros ;apply P_id_U71_bounded;assumption.
        intros ;apply P_id_isList_bounded;assumption.
        intros ;apply P_id_i_bounded;assumption.
        intros ;apply P_id_U11_bounded;assumption.
        intros ;apply P_id_isQid_bounded;assumption.
        intros ;apply P_id_isNeList_bounded;assumption.
        intros ;apply P_id_ok_bounded;assumption.
        intros ;apply P_id_mark_bounded;assumption.
        intros ;apply P_id_isPal_bounded;assumption.
        intros ;apply P_id_U41_bounded;assumption.
        intros ;apply P_id_u_bounded;assumption.
        intros ;apply P_id_U21_bounded;assumption.
        intros ;apply P_id_a_bounded;assumption.
        intros ;apply P_id_U52_bounded;assumption.
        intros ;apply P_id____bounded;assumption.
        intros ;apply P_id_U72_bounded;assumption.
        intros ;apply P_id_U31_bounded;assumption.
        intros ;apply P_id_o_bounded;assumption.
        intros ;apply P_id_tt_bounded;assumption.
        intros ;apply P_id_isNePal_bounded;assumption.
        intros ;apply P_id_U51_bounded;assumption.
        intros ;apply P_id_top_bounded;assumption.
        intros ;apply P_id_nil_bounded;assumption.
        intros ;apply P_id_U81_bounded;assumption.
        intros ;apply P_id_U42_bounded;assumption.
        intros ;apply P_id_proper_bounded;assumption.
        intros ;apply P_id_U22_bounded;assumption.
        intros ;apply P_id_e_bounded;assumption.
        intros ;apply P_id_U61_bounded;assumption.
        apply rules_monotonic.
        intros ;apply P_id_U22_hat_1_monotonic;assumption.
        intros ;apply P_id_ISNEPAL_monotonic;assumption.
        intros ;apply P_id_U21_hat_1_monotonic;assumption.
        intros ;apply P_id_U52_hat_1_monotonic;assumption.
        intros ;apply P_id_U51_hat_1_monotonic;assumption.
        intros ;apply P_id_U42_hat_1_monotonic;assumption.
        intros ;apply P_id_TOP_monotonic;assumption.
        intros ;apply P_id_ISQID_monotonic;assumption.
        intros ;apply P_id_ISPAL_monotonic;assumption.
        intros ;apply P_id_U71_hat_1_monotonic;assumption.
        intros ;apply P_id_ACTIVE_monotonic;assumption.
        intros ;apply P_id_ISLIST_monotonic;assumption.
        intros ;apply P_id_PROPER_monotonic;assumption.
        intros ;apply P_id_U31_hat_1_monotonic;assumption.
        intros ;apply P_id_U72_hat_1_monotonic;assumption.
        intros ;apply P_id_U61_hat_1_monotonic;assumption.
        intros ;apply P_id_ISNELIST_monotonic;assumption.
        intros ;apply P_id_U41_hat_1_monotonic;assumption.
        intros ;apply P_id_U11_hat_1_monotonic;assumption.
        intros ;apply P_id_U81_hat_1_monotonic;assumption.
        intros ;apply P_id____hat_1_monotonic;assumption.
      Qed.
      
      Ltac rewrite_and_unfold  :=
       do 2 (rewrite marked_measure_equation);
        repeat (
        match goal with
          |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
           rewrite (measure_equation (algebra.Alg.Term f t))
          | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
           rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
          end
        ).
      
      
      Lemma wf :
       well_founded WF_DP_R_xml_0_scc_37_large_large.DP_R_xml_0_scc_37_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_37_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 (x12:Z) := 2* x12.
     
     Definition P_id_U71 (x12:Z) (x13:Z) := 3.
     
     Definition P_id_isList (x12:Z) := 3.
     
     Definition P_id_i  := 2.
     
     Definition P_id_U11 (x12:Z) := 3* x12.
     
     Definition P_id_isQid (x12:Z) := 2 + 2* x12.
     
     Definition P_id_isNeList (x12:Z) := 2 + 3* x12.
     
     Definition P_id_ok (x12:Z) := 3.
     
     Definition P_id_mark (x12:Z) := 3.
     
     Definition P_id_isPal (x12:Z) := 2 + 1* x12.
     
     Definition P_id_U41 (x12:Z) (x13:Z) := 1 + 3* x12 + 2* x13.
     
     Definition P_id_u  := 2.
     
     Definition P_id_U21 (x12:Z) (x13:Z) := 3.
     
     Definition P_id_a  := 2.
     
     Definition P_id_U52 (x12:Z) := 2 + 2* x12.
     
     Definition P_id___ (x12:Z) (x13:Z) := 3 + 2* x12 + 2* x13.
     
     Definition P_id_U72 (x12:Z) := 2 + 1* x12.
     
     Definition P_id_U31 (x12:Z) := 2 + 2* x12.
     
     Definition P_id_o  := 2.
     
     Definition P_id_tt  := 1.
     
     Definition P_id_isNePal (x12:Z) := 2 + 2* x12.
     
     Definition P_id_U51 (x12:Z) (x13:Z) := 1 + 3* x12 + 2* x13.
     
     Definition P_id_top (x12:Z) := 0.
     
     Definition P_id_nil  := 2.
     
     Definition P_id_U81 (x12:Z) := 2 + 2* x12.
     
     Definition P_id_U42 (x12:Z) := 2 + 2* x12.
     
     Definition P_id_proper (x12:Z) := 3* x12.
     
     Definition P_id_U22 (x12:Z) := 2 + 2* x12.
     
     Definition P_id_e  := 2.
     
     Definition P_id_U61 (x12:Z) := 2 + 2* x12.
     
     Lemma P_id_active_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U71_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       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_isList_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U11_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isQid_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isNeList_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_ok_monotonic :
      forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
     Proof.
       intros x13 x12.
       intros [H_1 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 x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isPal_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U41_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       intros [H_1 H_0].
       intros [H_3 H_2].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U21_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       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_U52_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id____monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       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_U72_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U31_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isNePal_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U51_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       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_top_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U81_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U42_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_proper_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U22_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U61_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
     Proof.
       intros x13 x12.
       intros [H_1 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_bounded :
      forall x12, (0 <= x12) ->0 <= P_id_active x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U71_bounded :
      forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isList_bounded :
      forall x12, (0 <= x12) ->0 <= P_id_isList x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_i_bounded : 0 <= P_id_i .
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isNeList_bounded :
      forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
     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 x12, (0 <= x12) ->0 <= P_id_mark x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U41_bounded :
      forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_u_bounded : 0 <= P_id_u .
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U21_bounded :
      forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_a_bounded : 0 <= P_id_a .
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id____bounded :
      forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_o_bounded : 0 <= P_id_o .
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_tt_bounded : 0 <= P_id_tt .
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_isNePal_bounded :
      forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U51_bounded :
      forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
     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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_proper_bounded :
      forall x12, (0 <= x12) ->0 <= P_id_proper x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_e_bounded : 0 <= P_id_e .
     Proof.
       intros .
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
     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_U71 P_id_isList P_id_i P_id_U11 
        P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
        P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
        P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
        P_id_proper P_id_U22 P_id_e P_id_U61.
     
     Lemma measure_equation :
      forall t, 
       measure t = match t with
                     | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                      P_id_active (measure x12)
                     | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                      P_id_U71 (measure x13) (measure x12)
                     | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                      P_id_isList (measure x12)
                     | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                     | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                      P_id_U11 (measure x12)
                     | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                      P_id_isQid (measure x12)
                     | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                      P_id_isNeList (measure x12)
                     | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                      P_id_ok (measure x12)
                     | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                      P_id_mark (measure x12)
                     | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                      P_id_isPal (measure x12)
                     | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                      P_id_U41 (measure x13) (measure x12)
                     | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                     | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                      P_id_U21 (measure x13) (measure x12)
                     | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                     | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                      P_id_U52 (measure x12)
                     | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                      P_id___ (measure x13) (measure x12)
                     | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                      P_id_U72 (measure x12)
                     | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                      P_id_U31 (measure x12)
                     | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                     | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                     | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                      P_id_isNePal (measure x12)
                     | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                      P_id_U51 (measure x13) (measure x12)
                     | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                      P_id_top (measure x12)
                     | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                     | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                      P_id_U81 (measure x12)
                     | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                      P_id_U42 (measure x12)
                     | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                      P_id_proper (measure x12)
                     | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                      P_id_U22 (measure x12)
                     | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                     | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                      P_id_U61 (measure x12)
                     | _ => 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_U71_monotonic;assumption.
       intros ;apply P_id_isList_monotonic;assumption.
       intros ;apply P_id_U11_monotonic;assumption.
       intros ;apply P_id_isQid_monotonic;assumption.
       intros ;apply P_id_isNeList_monotonic;assumption.
       intros ;apply P_id_ok_monotonic;assumption.
       intros ;apply P_id_mark_monotonic;assumption.
       intros ;apply P_id_isPal_monotonic;assumption.
       intros ;apply P_id_U41_monotonic;assumption.
       intros ;apply P_id_U21_monotonic;assumption.
       intros ;apply P_id_U52_monotonic;assumption.
       intros ;apply P_id____monotonic;assumption.
       intros ;apply P_id_U72_monotonic;assumption.
       intros ;apply P_id_U31_monotonic;assumption.
       intros ;apply P_id_isNePal_monotonic;assumption.
       intros ;apply P_id_U51_monotonic;assumption.
       intros ;apply P_id_top_monotonic;assumption.
       intros ;apply P_id_U81_monotonic;assumption.
       intros ;apply P_id_U42_monotonic;assumption.
       intros ;apply P_id_proper_monotonic;assumption.
       intros ;apply P_id_U22_monotonic;assumption.
       intros ;apply P_id_U61_monotonic;assumption.
       intros ;apply P_id_active_bounded;assumption.
       intros ;apply P_id_U71_bounded;assumption.
       intros ;apply P_id_isList_bounded;assumption.
       intros ;apply P_id_i_bounded;assumption.
       intros ;apply P_id_U11_bounded;assumption.
       intros ;apply P_id_isQid_bounded;assumption.
       intros ;apply P_id_isNeList_bounded;assumption.
       intros ;apply P_id_ok_bounded;assumption.
       intros ;apply P_id_mark_bounded;assumption.
       intros ;apply P_id_isPal_bounded;assumption.
       intros ;apply P_id_U41_bounded;assumption.
       intros ;apply P_id_u_bounded;assumption.
       intros ;apply P_id_U21_bounded;assumption.
       intros ;apply P_id_a_bounded;assumption.
       intros ;apply P_id_U52_bounded;assumption.
       intros ;apply P_id____bounded;assumption.
       intros ;apply P_id_U72_bounded;assumption.
       intros ;apply P_id_U31_bounded;assumption.
       intros ;apply P_id_o_bounded;assumption.
       intros ;apply P_id_tt_bounded;assumption.
       intros ;apply P_id_isNePal_bounded;assumption.
       intros ;apply P_id_U51_bounded;assumption.
       intros ;apply P_id_top_bounded;assumption.
       intros ;apply P_id_nil_bounded;assumption.
       intros ;apply P_id_U81_bounded;assumption.
       intros ;apply P_id_U42_bounded;assumption.
       intros ;apply P_id_proper_bounded;assumption.
       intros ;apply P_id_U22_bounded;assumption.
       intros ;apply P_id_e_bounded;assumption.
       intros ;apply P_id_U61_bounded;assumption.
       apply rules_monotonic.
     Qed.
     
     Definition P_id_U22_hat_1 (x12:Z) := 0.
     
     Definition P_id_ISNEPAL (x12:Z) := 0.
     
     Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
     
     Definition P_id_U52_hat_1 (x12:Z) := 0.
     
     Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
     
     Definition P_id_U42_hat_1 (x12:Z) := 0.
     
     Definition P_id_TOP (x12:Z) := 0.
     
     Definition P_id_ISQID (x12:Z) := 0.
     
     Definition P_id_ISPAL (x12:Z) := 0.
     
     Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
     
     Definition P_id_ACTIVE (x12:Z) := 0.
     
     Definition P_id_ISLIST (x12:Z) := 0.
     
     Definition P_id_PROPER (x12:Z) := 1* x12.
     
     Definition P_id_U31_hat_1 (x12:Z) := 0.
     
     Definition P_id_U72_hat_1 (x12:Z) := 0.
     
     Definition P_id_U61_hat_1 (x12:Z) := 0.
     
     Definition P_id_ISNELIST (x12:Z) := 0.
     
     Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
     
     Definition P_id_U11_hat_1 (x12:Z) := 0.
     
     Definition P_id_U81_hat_1 (x12:Z) := 0.
     
     Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
     
     Lemma P_id_U22_hat_1_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_ISNEPAL_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U21_hat_1_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->
         P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       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_U52_hat_1_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U51_hat_1_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->
         P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       intros [H_1 H_0].
       intros [H_3 H_2].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U42_hat_1_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_TOP_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_ISQID_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_ISPAL_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U71_hat_1_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->
         P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       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_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_ISLIST_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_PROPER_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U31_hat_1_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U72_hat_1_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U61_hat_1_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_ISNELIST_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U41_hat_1_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->
         P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       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_U11_hat_1_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id_U81_hat_1_monotonic :
      forall x12 x13, 
       (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
     Proof.
       intros x13 x12.
       intros [H_1 H_0].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Lemma P_id____hat_1_monotonic :
      forall x12 x14 x13 x15, 
       (0 <= x15)/\ (x15 <= x14) ->
        (0 <= x13)/\ (x13 <= x12) ->
         P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
     Proof.
       intros x15 x14 x13 x12.
       intros [H_1 H_0].
       intros [H_3 H_2].
       
       cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
        (auto with zarith)||(repeat (translate_vars );prove_ineq ).
     Qed.
     
     Definition marked_measure  := 
       InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
        P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
        P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
        P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
        P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
        P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
        P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
        P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
        P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
        P_id_U81_hat_1 P_id____hat_1.
     
     Lemma marked_measure_equation :
      forall t, 
       marked_measure t = match t with
                            | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                             P_id_U22_hat_1 (measure x12)
                            | (algebra.Alg.Term algebra.F.id_isNePal 
                               (x12::nil)) =>
                             P_id_ISNEPAL (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U21 (x13::
                               x12::nil)) =>
                             P_id_U21_hat_1 (measure x13) (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                             P_id_U52_hat_1 (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U51 (x13::
                               x12::nil)) =>
                             P_id_U51_hat_1 (measure x13) (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                             P_id_U42_hat_1 (measure x12)
                            | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                             P_id_TOP (measure x12)
                            | (algebra.Alg.Term algebra.F.id_isQid 
                               (x12::nil)) =>
                             P_id_ISQID (measure x12)
                            | (algebra.Alg.Term algebra.F.id_isPal 
                               (x12::nil)) =>
                             P_id_ISPAL (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U71 (x13::
                               x12::nil)) =>
                             P_id_U71_hat_1 (measure x13) (measure x12)
                            | (algebra.Alg.Term algebra.F.id_active 
                               (x12::nil)) =>
                             P_id_ACTIVE (measure x12)
                            | (algebra.Alg.Term algebra.F.id_isList 
                               (x12::nil)) =>
                             P_id_ISLIST (measure x12)
                            | (algebra.Alg.Term algebra.F.id_proper 
                               (x12::nil)) =>
                             P_id_PROPER (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                             P_id_U31_hat_1 (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                             P_id_U72_hat_1 (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                             P_id_U61_hat_1 (measure x12)
                            | (algebra.Alg.Term algebra.F.id_isNeList 
                               (x12::nil)) =>
                             P_id_ISNELIST (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U41 (x13::
                               x12::nil)) =>
                             P_id_U41_hat_1 (measure x13) (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                             P_id_U11_hat_1 (measure x12)
                            | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                             P_id_U81_hat_1 (measure x12)
                            | (algebra.Alg.Term algebra.F.id___ (x13::
                               x12::nil)) =>
                             P_id____hat_1 (measure x13) (measure x12)
                            | _ => 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_U71_monotonic;assumption.
       intros ;apply P_id_isList_monotonic;assumption.
       intros ;apply P_id_U11_monotonic;assumption.
       intros ;apply P_id_isQid_monotonic;assumption.
       intros ;apply P_id_isNeList_monotonic;assumption.
       intros ;apply P_id_ok_monotonic;assumption.
       intros ;apply P_id_mark_monotonic;assumption.
       intros ;apply P_id_isPal_monotonic;assumption.
       intros ;apply P_id_U41_monotonic;assumption.
       intros ;apply P_id_U21_monotonic;assumption.
       intros ;apply P_id_U52_monotonic;assumption.
       intros ;apply P_id____monotonic;assumption.
       intros ;apply P_id_U72_monotonic;assumption.
       intros ;apply P_id_U31_monotonic;assumption.
       intros ;apply P_id_isNePal_monotonic;assumption.
       intros ;apply P_id_U51_monotonic;assumption.
       intros ;apply P_id_top_monotonic;assumption.
       intros ;apply P_id_U81_monotonic;assumption.
       intros ;apply P_id_U42_monotonic;assumption.
       intros ;apply P_id_proper_monotonic;assumption.
       intros ;apply P_id_U22_monotonic;assumption.
       intros ;apply P_id_U61_monotonic;assumption.
       intros ;apply P_id_active_bounded;assumption.
       intros ;apply P_id_U71_bounded;assumption.
       intros ;apply P_id_isList_bounded;assumption.
       intros ;apply P_id_i_bounded;assumption.
       intros ;apply P_id_U11_bounded;assumption.
       intros ;apply P_id_isQid_bounded;assumption.
       intros ;apply P_id_isNeList_bounded;assumption.
       intros ;apply P_id_ok_bounded;assumption.
       intros ;apply P_id_mark_bounded;assumption.
       intros ;apply P_id_isPal_bounded;assumption.
       intros ;apply P_id_U41_bounded;assumption.
       intros ;apply P_id_u_bounded;assumption.
       intros ;apply P_id_U21_bounded;assumption.
       intros ;apply P_id_a_bounded;assumption.
       intros ;apply P_id_U52_bounded;assumption.
       intros ;apply P_id____bounded;assumption.
       intros ;apply P_id_U72_bounded;assumption.
       intros ;apply P_id_U31_bounded;assumption.
       intros ;apply P_id_o_bounded;assumption.
       intros ;apply P_id_tt_bounded;assumption.
       intros ;apply P_id_isNePal_bounded;assumption.
       intros ;apply P_id_U51_bounded;assumption.
       intros ;apply P_id_top_bounded;assumption.
       intros ;apply P_id_nil_bounded;assumption.
       intros ;apply P_id_U81_bounded;assumption.
       intros ;apply P_id_U42_bounded;assumption.
       intros ;apply P_id_proper_bounded;assumption.
       intros ;apply P_id_U22_bounded;assumption.
       intros ;apply P_id_e_bounded;assumption.
       intros ;apply P_id_U61_bounded;assumption.
       apply rules_monotonic.
       intros ;apply P_id_U22_hat_1_monotonic;assumption.
       intros ;apply P_id_ISNEPAL_monotonic;assumption.
       intros ;apply P_id_U21_hat_1_monotonic;assumption.
       intros ;apply P_id_U52_hat_1_monotonic;assumption.
       intros ;apply P_id_U51_hat_1_monotonic;assumption.
       intros ;apply P_id_U42_hat_1_monotonic;assumption.
       intros ;apply P_id_TOP_monotonic;assumption.
       intros ;apply P_id_ISQID_monotonic;assumption.
       intros ;apply P_id_ISPAL_monotonic;assumption.
       intros ;apply P_id_U71_hat_1_monotonic;assumption.
       intros ;apply P_id_ACTIVE_monotonic;assumption.
       intros ;apply P_id_ISLIST_monotonic;assumption.
       intros ;apply P_id_PROPER_monotonic;assumption.
       intros ;apply P_id_U31_hat_1_monotonic;assumption.
       intros ;apply P_id_U72_hat_1_monotonic;assumption.
       intros ;apply P_id_U61_hat_1_monotonic;assumption.
       intros ;apply P_id_ISNELIST_monotonic;assumption.
       intros ;apply P_id_U41_hat_1_monotonic;assumption.
       intros ;apply P_id_U11_hat_1_monotonic;assumption.
       intros ;apply P_id_U81_hat_1_monotonic;assumption.
       intros ;apply P_id____hat_1_monotonic;assumption.
     Qed.
     
     Ltac rewrite_and_unfold  :=
      do 2 (rewrite marked_measure_equation);
       repeat (
       match goal with
         |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
          rewrite (measure_equation (algebra.Alg.Term f t))
         | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
          rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
         end
       ).
     
     Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
     
     Definition le a b := marked_measure a <= marked_measure b.
     
     Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
     Proof.
       unfold lt, le in *.
       intros a b c.
       apply (interp.le_lt_compat_right (interp.o_Z 0)).
     Qed.
     
     Lemma wf_lt : well_founded lt.
     Proof.
       unfold lt in *.
       apply Inverse_Image.wf_inverse_image with  (B:=Z).
       apply Zwf.Zwf_well_founded.
     Qed.
     
     Lemma DP_R_xml_0_scc_37_large_large_strict_in_lt :
      Relation_Definitions.inclusion _ DP_R_xml_0_scc_37_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_37_large_large_large_in_le :
      Relation_Definitions.inclusion _ DP_R_xml_0_scc_37_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_37_large_large_large  := 
       WF_DP_R_xml_0_scc_37_large_large_large.wf.
     
     
     Lemma wf :
      well_founded WF_DP_R_xml_0_scc_37_large.DP_R_xml_0_scc_37_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_37_large_large_large).
       clear x.
       intros x _ IHx IHx'.
       constructor.
       intros y H.
       
       destruct H;
        (apply IHx';apply DP_R_xml_0_scc_37_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_37_large_large_large_in_le;
           econstructor eassumption])).
       apply wf_DP_R_xml_0_scc_37_large_large_large.
     Qed.
    End WF_DP_R_xml_0_scc_37_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 (x12:Z) := 2* x12.
    
    Definition P_id_U71 (x12:Z) (x13:Z) := 3* x13.
    
    Definition P_id_isList (x12:Z) := 3 + 2* x12.
    
    Definition P_id_i  := 2.
    
    Definition P_id_U11 (x12:Z) := 2* x12.
    
    Definition P_id_isQid (x12:Z) := 1 + 2* x12.
    
    Definition P_id_isNeList (x12:Z) := 3 + 2* x12.
    
    Definition P_id_ok (x12:Z) := 0.
    
    Definition P_id_mark (x12:Z) := 0.
    
    Definition P_id_isPal (x12:Z) := 2 + 1* x12.
    
    Definition P_id_U41 (x12:Z) (x13:Z) := 3* x13.
    
    Definition P_id_u  := 2.
    
    Definition P_id_U21 (x12:Z) (x13:Z) := 1 + 1* x12 + 2* x13.
    
    Definition P_id_a  := 2.
    
    Definition P_id_U52 (x12:Z) := 3.
    
    Definition P_id___ (x12:Z) (x13:Z) := 2* x12 + 2* x13.
    
    Definition P_id_U72 (x12:Z) := 0.
    
    Definition P_id_U31 (x12:Z) := 2.
    
    Definition P_id_o  := 2.
    
    Definition P_id_tt  := 0.
    
    Definition P_id_isNePal (x12:Z) := 2 + 1* x12.
    
    Definition P_id_U51 (x12:Z) (x13:Z) := 3.
    
    Definition P_id_top (x12:Z) := 0.
    
    Definition P_id_nil  := 1.
    
    Definition P_id_U81 (x12:Z) := 1.
    
    Definition P_id_U42 (x12:Z) := 3.
    
    Definition P_id_proper (x12:Z) := 2* x12.
    
    Definition P_id_U22 (x12:Z) := 3.
    
    Definition P_id_e  := 2.
    
    Definition P_id_U61 (x12:Z) := 0.
    
    Lemma P_id_active_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_isList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
    Proof.
      intros x13 x12.
      intros [H_1 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 x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U72_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_top_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
    Proof.
      intros x13 x12.
      intros [H_1 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_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_active x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_i_bounded : 0 <= P_id_i .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
    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 x12, (0 <= x12) ->0 <= P_id_mark x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_u_bounded : 0 <= P_id_u .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_a_bounded : 0 <= P_id_a .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_o_bounded : 0 <= P_id_o .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_tt_bounded : 0 <= P_id_tt .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
    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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_proper x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_e_bounded : 0 <= P_id_e .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
    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_U71 P_id_isList P_id_i P_id_U11 
       P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
       P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
       P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
       P_id_proper P_id_U22 P_id_e P_id_U61.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                     P_id_active (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                     P_id_U71 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                     P_id_isList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                    | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                     P_id_U11 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                     P_id_isQid (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                     P_id_isNeList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                     P_id_ok (measure x12)
                    | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                     P_id_mark (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                     P_id_isPal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                     P_id_U41 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                    | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                     P_id_U21 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                    | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                     P_id_U52 (measure x12)
                    | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                     P_id___ (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                     P_id_U72 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                     P_id_U31 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                    | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                    | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                     P_id_isNePal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                     P_id_U51 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                     P_id_top (measure x12)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                     P_id_U81 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                     P_id_U42 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                     P_id_proper (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                     P_id_U22 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                    | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                     P_id_U61 (measure x12)
                    | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_U22_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNEPAL (x12:Z) := 0.
    
    Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U52_hat_1 (x12:Z) := 0.
    
    Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U42_hat_1 (x12:Z) := 0.
    
    Definition P_id_TOP (x12:Z) := 0.
    
    Definition P_id_ISQID (x12:Z) := 0.
    
    Definition P_id_ISPAL (x12:Z) := 0.
    
    Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_ACTIVE (x12:Z) := 0.
    
    Definition P_id_ISLIST (x12:Z) := 0.
    
    Definition P_id_PROPER (x12:Z) := 1* x12.
    
    Definition P_id_U31_hat_1 (x12:Z) := 0.
    
    Definition P_id_U72_hat_1 (x12:Z) := 0.
    
    Definition P_id_U61_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNELIST (x12:Z) := 0.
    
    Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U11_hat_1 (x12:Z) := 0.
    
    Definition P_id_U81_hat_1 (x12:Z) := 0.
    
    Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
    
    Lemma P_id_U22_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNEPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_TOP_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISQID_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISLIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_PROPER_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNELIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U11_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Definition marked_measure  := 
      InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
       P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
       P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
       P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
       P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
       P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
       P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
       P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
       P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
       P_id_U81_hat_1 P_id____hat_1.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                            P_id_U22_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNePal 
                              (x12::nil)) =>
                            P_id_ISNEPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U21 (x13::
                              x12::nil)) =>
                            P_id_U21_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                            P_id_U52_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U51 (x13::
                              x12::nil)) =>
                            P_id_U51_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                            P_id_U42_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                            P_id_TOP (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                            P_id_ISQID (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                            P_id_ISPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U71 (x13::
                              x12::nil)) =>
                            P_id_U71_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_active 
                              (x12::nil)) =>
                            P_id_ACTIVE (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isList 
                              (x12::nil)) =>
                            P_id_ISLIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_proper 
                              (x12::nil)) =>
                            P_id_PROPER (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                            P_id_U31_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                            P_id_U72_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                            P_id_U61_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNeList 
                              (x12::nil)) =>
                            P_id_ISNELIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U41 (x13::
                              x12::nil)) =>
                            P_id_U41_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                            P_id_U11_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                            P_id_U81_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id___ (x13::
                              x12::nil)) =>
                            P_id____hat_1 (measure x13) (measure x12)
                           | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_U22_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNEPAL_monotonic;assumption.
      intros ;apply P_id_U21_hat_1_monotonic;assumption.
      intros ;apply P_id_U52_hat_1_monotonic;assumption.
      intros ;apply P_id_U51_hat_1_monotonic;assumption.
      intros ;apply P_id_U42_hat_1_monotonic;assumption.
      intros ;apply P_id_TOP_monotonic;assumption.
      intros ;apply P_id_ISQID_monotonic;assumption.
      intros ;apply P_id_ISPAL_monotonic;assumption.
      intros ;apply P_id_U71_hat_1_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_ISLIST_monotonic;assumption.
      intros ;apply P_id_PROPER_monotonic;assumption.
      intros ;apply P_id_U31_hat_1_monotonic;assumption.
      intros ;apply P_id_U72_hat_1_monotonic;assumption.
      intros ;apply P_id_U61_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNELIST_monotonic;assumption.
      intros ;apply P_id_U41_hat_1_monotonic;assumption.
      intros ;apply P_id_U11_hat_1_monotonic;assumption.
      intros ;apply P_id_U81_hat_1_monotonic;assumption.
      intros ;apply P_id____hat_1_monotonic;assumption.
    Qed.
    
    Ltac rewrite_and_unfold  :=
     do 2 (rewrite marked_measure_equation);
      repeat (
      match goal with
        |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
         rewrite (measure_equation (algebra.Alg.Term f t))
        | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
         rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
        end
      ).
    
    Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
    
    Definition le a b := marked_measure a <= marked_measure b.
    
    Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
    Proof.
      unfold lt, le in *.
      intros a b c.
      apply (interp.le_lt_compat_right (interp.o_Z 0)).
    Qed.
    
    Lemma wf_lt : well_founded lt.
    Proof.
      unfold lt in *.
      apply Inverse_Image.wf_inverse_image with  (B:=Z).
      apply Zwf.Zwf_well_founded.
    Qed.
    
    Lemma DP_R_xml_0_scc_37_large_strict_in_lt :
     Relation_Definitions.inclusion _ DP_R_xml_0_scc_37_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_37_large_large_in_le :
     Relation_Definitions.inclusion _ DP_R_xml_0_scc_37_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_37_large_large  := 
      WF_DP_R_xml_0_scc_37_large_large.wf.
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_37.DP_R_xml_0_scc_37_large.
    Proof.
      intros x.
      apply (well_founded_ind wf_lt).
      clear x.
      intros x.
      pattern x.
      apply (@Acc_ind _ DP_R_xml_0_scc_37_large_large).
      clear x.
      intros x _ IHx IHx'.
      constructor.
      intros y H.
      
      destruct H;
       (apply IHx';apply DP_R_xml_0_scc_37_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_37_large_large_in_le;econstructor eassumption])).
      apply wf_DP_R_xml_0_scc_37_large_large.
    Qed.
   End WF_DP_R_xml_0_scc_37_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 (x12:Z) := 2* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 2 + 2* x12 + 1* x13.
   
   Definition P_id_isList (x12:Z) := 2* x12.
   
   Definition P_id_i  := 2.
   
   Definition P_id_U11 (x12:Z) := 1* x12.
   
   Definition P_id_isQid (x12:Z) := 1 + 1* x12.
   
   Definition P_id_isNeList (x12:Z) := 1 + 2* x12.
   
   Definition P_id_ok (x12:Z) := 0.
   
   Definition P_id_mark (x12:Z) := 0.
   
   Definition P_id_isPal (x12:Z) := 3 + 1* x12.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 1 + 1* x12 + 2* x13.
   
   Definition P_id_u  := 0.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 2* x12 + 3* x13.
   
   Definition P_id_a  := 2.
   
   Definition P_id_U52 (x12:Z) := 3 + 2* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 1* x12 + 2* x13.
   
   Definition P_id_U72 (x12:Z) := 1 + 1* x12.
   
   Definition P_id_U31 (x12:Z) := 3 + 2* x12.
   
   Definition P_id_o  := 2.
   
   Definition P_id_tt  := 0.
   
   Definition P_id_isNePal (x12:Z) := 3 + 1* x12.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 3 + 1* x12 + 1* x13.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 2.
   
   Definition P_id_U81 (x12:Z) := 3 + 1* x12.
   
   Definition P_id_U42 (x12:Z) := 3 + 2* x12.
   
   Definition P_id_proper (x12:Z) := 2* x12.
   
   Definition P_id_U22 (x12:Z) := 3 + 2* x12.
   
   Definition P_id_e  := 2.
   
   Definition P_id_U61 (x12:Z) := 3 + 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 2* x12.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
   
   Definition le a b := marked_measure a <= marked_measure b.
   
   Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
   Proof.
     unfold lt, le in *.
     intros a b c.
     apply (interp.le_lt_compat_right (interp.o_Z 0)).
   Qed.
   
   Lemma wf_lt : well_founded lt.
   Proof.
     unfold lt in *.
     apply Inverse_Image.wf_inverse_image with  (B:=Z).
     apply Zwf.Zwf_well_founded.
   Qed.
   
   Lemma DP_R_xml_0_scc_37_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_37_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_37_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_37_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_37_large  := WF_DP_R_xml_0_scc_37_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_37.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_37_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_37_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_37_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_37_large.
   Qed.
  End WF_DP_R_xml_0_scc_37.
  
  Definition wf_DP_R_xml_0_scc_37  := WF_DP_R_xml_0_scc_37.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_37 :
   forall x y, (DP_R_xml_0_scc_37 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_37).
    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_36;
         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_32;
           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_28;
             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_24;
               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_20;
                 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_16;
                   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_12;
                     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' ))||
                        ((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_37.
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_38  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <top(mark(X_)),proper(X_)> *)
    | DP_R_xml_0_non_scc_38_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_38 (algebra.Alg.Term algebra.F.id_proper (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_top (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_38 :
   forall x y, 
    (DP_R_xml_0_non_scc_38 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_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_34;
          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_30;
            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_26;
              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_22;
                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_18;
                  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_14;
                    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_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' ))||
                        ((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_39  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U81(X_)),U81(active(X_))> *)
    | DP_R_xml_0_non_scc_39_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U81 (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_39 (algebra.Alg.Term algebra.F.id_U81 
                              ((algebra.Alg.Term algebra.F.id_active 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_39 :
   forall x y, 
    (DP_R_xml_0_non_scc_39 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_11;
        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_40  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U72(X_)),U72(active(X_))> *)
    | DP_R_xml_0_non_scc_40_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_40 (algebra.Alg.Term algebra.F.id_U72 
                              ((algebra.Alg.Term algebra.F.id_active 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_40 :
   forall x y, 
    (DP_R_xml_0_non_scc_40 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_13;
        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_41  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U71(X1_,X2_)),U71(active(X1_),X2_)> *)
    | DP_R_xml_0_non_scc_41_0 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x12) ->
       DP_R_xml_0_non_scc_41 (algebra.Alg.Term algebra.F.id_U71 
                              ((algebra.Alg.Term algebra.F.id_active 
                              (x9::nil))::x10::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_41 :
   forall x y, 
    (DP_R_xml_0_non_scc_41 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_15;
        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_42  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U61(X_)),U61(active(X_))> *)
    | DP_R_xml_0_non_scc_42_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U61 (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_42 (algebra.Alg.Term algebra.F.id_U61 
                              ((algebra.Alg.Term algebra.F.id_active 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_42 :
   forall x y, 
    (DP_R_xml_0_non_scc_42 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_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_43  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U52(X_)),U52(active(X_))> *)
    | DP_R_xml_0_non_scc_43_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U52 (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_43 (algebra.Alg.Term algebra.F.id_U52 
                              ((algebra.Alg.Term algebra.F.id_active 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_43 :
   forall x y, 
    (DP_R_xml_0_non_scc_43 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_19;
        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_44  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U51(X1_,X2_)),U51(active(X1_),X2_)> *)
    | DP_R_xml_0_non_scc_44_0 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U51 (x9::x10::nil)) x12) ->
       DP_R_xml_0_non_scc_44 (algebra.Alg.Term algebra.F.id_U51 
                              ((algebra.Alg.Term algebra.F.id_active 
                              (x9::nil))::x10::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_44 :
   forall x y, 
    (DP_R_xml_0_non_scc_44 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_21;
        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_45  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U42(X_)),U42(active(X_))> *)
    | DP_R_xml_0_non_scc_45_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U42 (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_45 (algebra.Alg.Term algebra.F.id_U42 
                              ((algebra.Alg.Term algebra.F.id_active 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_45 :
   forall x y, 
    (DP_R_xml_0_non_scc_45 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_23;
        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_46  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U41(X1_,X2_)),U41(active(X1_),X2_)> *)
    | DP_R_xml_0_non_scc_46_0 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U41 (x9::x10::nil)) x12) ->
       DP_R_xml_0_non_scc_46 (algebra.Alg.Term algebra.F.id_U41 
                              ((algebra.Alg.Term algebra.F.id_active 
                              (x9::nil))::x10::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_46 :
   forall x y, 
    (DP_R_xml_0_non_scc_46 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_25;
        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_47  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U31(X_)),U31(active(X_))> *)
    | DP_R_xml_0_non_scc_47_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U31 (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_47 (algebra.Alg.Term algebra.F.id_U31 
                              ((algebra.Alg.Term algebra.F.id_active 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_47 :
   forall x y, 
    (DP_R_xml_0_non_scc_47 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_27;
        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_48  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U22(X_)),U22(active(X_))> *)
    | DP_R_xml_0_non_scc_48_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U22 (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_48 (algebra.Alg.Term algebra.F.id_U22 
                              ((algebra.Alg.Term algebra.F.id_active 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_48 :
   forall x y, 
    (DP_R_xml_0_non_scc_48 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_29;
        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_49  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U21(X1_,X2_)),U21(active(X1_),X2_)> *)
    | DP_R_xml_0_non_scc_49_0 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) x12) ->
       DP_R_xml_0_non_scc_49 (algebra.Alg.Term algebra.F.id_U21 
                              ((algebra.Alg.Term algebra.F.id_active 
                              (x9::nil))::x10::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_49 :
   forall x y, 
    (DP_R_xml_0_non_scc_49 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_31;
        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_50  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U11(X_)),U11(active(X_))> *)
    | DP_R_xml_0_non_scc_50_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U11 (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_50 (algebra.Alg.Term algebra.F.id_U11 
                              ((algebra.Alg.Term algebra.F.id_active 
                              (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_50 :
   forall x y, 
    (DP_R_xml_0_non_scc_50 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_33;
        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_51  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(__(X1_,X2_)),__(X1_,active(X2_))> *)
    | DP_R_xml_0_non_scc_51_0 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
       DP_R_xml_0_non_scc_51 (algebra.Alg.Term algebra.F.id___ (x9::
                              (algebra.Alg.Term algebra.F.id_active 
                              (x10::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_51 :
   forall x y, 
    (DP_R_xml_0_non_scc_51 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_35;
        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_52  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(__(X1_,X2_)),__(active(X1_),X2_)> *)
    | DP_R_xml_0_non_scc_52_0 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
       DP_R_xml_0_non_scc_52 (algebra.Alg.Term algebra.F.id___ 
                              ((algebra.Alg.Term algebra.F.id_active 
                              (x9::nil))::x10::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_52 :
   forall x y, 
    (DP_R_xml_0_non_scc_52 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_35;
        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_53  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(isPal(V_)),isNePal(V_)> *)
    | DP_R_xml_0_non_scc_53_0 :
     forall x12 x6, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isPal (x6::nil)) x12) ->
       DP_R_xml_0_non_scc_53 (algebra.Alg.Term algebra.F.id_isNePal 
                              (x6::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_53 :
   forall x y, 
    (DP_R_xml_0_non_scc_53 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_54  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(isPal(V_)),U81(isNePal(V_))> *)
    | DP_R_xml_0_non_scc_54_0 :
     forall x12 x6, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isPal (x6::nil)) x12) ->
       DP_R_xml_0_non_scc_54 (algebra.Alg.Term algebra.F.id_U81 
                              ((algebra.Alg.Term algebra.F.id_isNePal 
                              (x6::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_54 :
   forall x y, 
    (DP_R_xml_0_non_scc_54 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_11;
        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_55  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(isNePal(__(I_,__(P_,I_)))),isQid(I_)> *)
    | DP_R_xml_0_non_scc_55_0 :
     forall x8 x12 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isNePal ((algebra.Alg.Term 
         algebra.F.id___ (x8::(algebra.Alg.Term algebra.F.id___ (x5::
         x8::nil))::nil))::nil)) x12) ->
       DP_R_xml_0_non_scc_55 (algebra.Alg.Term algebra.F.id_isQid (x8::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_55 :
   forall x y, 
    (DP_R_xml_0_non_scc_55 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_56  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(isNePal(__(I_,__(P_,I_)))),U71(isQid(I_),P_)> *)
    | DP_R_xml_0_non_scc_56_0 :
     forall x8 x12 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isNePal ((algebra.Alg.Term 
         algebra.F.id___ (x8::(algebra.Alg.Term algebra.F.id___ (x5::
         x8::nil))::nil))::nil)) x12) ->
       DP_R_xml_0_non_scc_56 (algebra.Alg.Term algebra.F.id_U71 
                              ((algebra.Alg.Term algebra.F.id_isQid 
                              (x8::nil))::x5::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_56 :
   forall x y, 
    (DP_R_xml_0_non_scc_56 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_15;
        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_57  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(isNePal(V_)),isQid(V_)> *)
    | DP_R_xml_0_non_scc_57_0 :
     forall x12 x6, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isNePal (x6::nil)) x12) ->
       DP_R_xml_0_non_scc_57 (algebra.Alg.Term algebra.F.id_isQid (x6::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_57 :
   forall x y, 
    (DP_R_xml_0_non_scc_57 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_58  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(isNePal(V_)),U61(isQid(V_))> *)
    | DP_R_xml_0_non_scc_58_0 :
     forall x12 x6, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isNePal (x6::nil)) x12) ->
       DP_R_xml_0_non_scc_58 (algebra.Alg.Term algebra.F.id_U61 
                              ((algebra.Alg.Term algebra.F.id_isQid 
                              (x6::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_58 :
   forall x y, 
    (DP_R_xml_0_non_scc_58 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_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_59  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(isNeList(__(V1_,V2_))),isNeList(V1_)> *)
    | DP_R_xml_0_non_scc_59_0 :
     forall x4 x12 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Term 
         algebra.F.id___ (x7::x4::nil))::nil)) x12) ->
       DP_R_xml_0_non_scc_59 (algebra.Alg.Term algebra.F.id_isNeList 
                              (x7::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_59 :
   forall x y, 
    (DP_R_xml_0_non_scc_59 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_60  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(isNeList(__(V1_,V2_))),U51(isNeList(V1_),V2_)> *)
    | DP_R_xml_0_non_scc_60_0 :
     forall x4 x12 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Term 
         algebra.F.id___ (x7::x4::nil))::nil)) x12) ->
       DP_R_xml_0_non_scc_60 (algebra.Alg.Term algebra.F.id_U51 
                              ((algebra.Alg.Term algebra.F.id_isNeList 
                              (x7::nil))::x4::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_60 :
   forall x y, 
    (DP_R_xml_0_non_scc_60 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_21;
        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_61  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(isNeList(__(V1_,V2_))),isList(V1_)> *)
    | DP_R_xml_0_non_scc_61_0 :
     forall x4 x12 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Term 
         algebra.F.id___ (x7::x4::nil))::nil)) x12) ->
       DP_R_xml_0_non_scc_61 (algebra.Alg.Term algebra.F.id_isList (x7::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_61 :
   forall x y, 
    (DP_R_xml_0_non_scc_61 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_62  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(isNeList(__(V1_,V2_))),U41(isList(V1_),V2_)> *)
    | DP_R_xml_0_non_scc_62_0 :
     forall x4 x12 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Term 
         algebra.F.id___ (x7::x4::nil))::nil)) x12) ->
       DP_R_xml_0_non_scc_62 (algebra.Alg.Term algebra.F.id_U41 
                              ((algebra.Alg.Term algebra.F.id_isList 
                              (x7::nil))::x4::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_62 :
   forall x y, 
    (DP_R_xml_0_non_scc_62 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_25;
        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_63  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(isNeList(V_)),isQid(V_)> *)
    | DP_R_xml_0_non_scc_63_0 :
     forall x12 x6, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isNeList (x6::nil)) x12) ->
       DP_R_xml_0_non_scc_63 (algebra.Alg.Term algebra.F.id_isQid (x6::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_63 :
   forall x y, 
    (DP_R_xml_0_non_scc_63 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_64  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(isNeList(V_)),U31(isQid(V_))> *)
    | DP_R_xml_0_non_scc_64_0 :
     forall x12 x6, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isNeList (x6::nil)) x12) ->
       DP_R_xml_0_non_scc_64 (algebra.Alg.Term algebra.F.id_U31 
                              ((algebra.Alg.Term algebra.F.id_isQid 
                              (x6::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_64 :
   forall x y, 
    (DP_R_xml_0_non_scc_64 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_27;
        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_65  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(isList(__(V1_,V2_))),isList(V1_)> *)
    | DP_R_xml_0_non_scc_65_0 :
     forall x4 x12 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isList ((algebra.Alg.Term 
         algebra.F.id___ (x7::x4::nil))::nil)) x12) ->
       DP_R_xml_0_non_scc_65 (algebra.Alg.Term algebra.F.id_isList (x7::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_65 :
   forall x y, 
    (DP_R_xml_0_non_scc_65 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_66  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(isList(__(V1_,V2_))),U21(isList(V1_),V2_)> *)
    | DP_R_xml_0_non_scc_66_0 :
     forall x4 x12 x7, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isList ((algebra.Alg.Term 
         algebra.F.id___ (x7::x4::nil))::nil)) x12) ->
       DP_R_xml_0_non_scc_66 (algebra.Alg.Term algebra.F.id_U21 
                              ((algebra.Alg.Term algebra.F.id_isList 
                              (x7::nil))::x4::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_66 :
   forall x y, 
    (DP_R_xml_0_non_scc_66 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_31;
        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_67  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(isList(V_)),isNeList(V_)> *)
    | DP_R_xml_0_non_scc_67_0 :
     forall x12 x6, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isList (x6::nil)) x12) ->
       DP_R_xml_0_non_scc_67 (algebra.Alg.Term algebra.F.id_isNeList 
                              (x6::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_67 :
   forall x y, 
    (DP_R_xml_0_non_scc_67 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_68  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(isList(V_)),U11(isNeList(V_))> *)
    | DP_R_xml_0_non_scc_68_0 :
     forall x12 x6, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_isList (x6::nil)) x12) ->
       DP_R_xml_0_non_scc_68 (algebra.Alg.Term algebra.F.id_U11 
                              ((algebra.Alg.Term algebra.F.id_isNeList 
                              (x6::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_68 :
   forall x y, 
    (DP_R_xml_0_non_scc_68 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_33;
        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_69  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U71(tt,P_)),isPal(P_)> *)
    | DP_R_xml_0_non_scc_69_0 :
     forall x12 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Term 
         algebra.F.id_tt nil)::x5::nil)) x12) ->
       DP_R_xml_0_non_scc_69 (algebra.Alg.Term algebra.F.id_isPal (x5::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_69 :
   forall x y, 
    (DP_R_xml_0_non_scc_69 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_70  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U71(tt,P_)),U72(isPal(P_))> *)
    | DP_R_xml_0_non_scc_70_0 :
     forall x12 x5, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Term 
         algebra.F.id_tt nil)::x5::nil)) x12) ->
       DP_R_xml_0_non_scc_70 (algebra.Alg.Term algebra.F.id_U72 
                              ((algebra.Alg.Term algebra.F.id_isPal 
                              (x5::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_70 :
   forall x y, 
    (DP_R_xml_0_non_scc_70 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_13;
        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_71  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U51(tt,V2_)),isList(V2_)> *)
    | DP_R_xml_0_non_scc_71_0 :
     forall x4 x12, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Term 
         algebra.F.id_tt nil)::x4::nil)) x12) ->
       DP_R_xml_0_non_scc_71 (algebra.Alg.Term algebra.F.id_isList (x4::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_71 :
   forall x y, 
    (DP_R_xml_0_non_scc_71 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_72  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U51(tt,V2_)),U52(isList(V2_))> *)
    | DP_R_xml_0_non_scc_72_0 :
     forall x4 x12, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Term 
         algebra.F.id_tt nil)::x4::nil)) x12) ->
       DP_R_xml_0_non_scc_72 (algebra.Alg.Term algebra.F.id_U52 
                              ((algebra.Alg.Term algebra.F.id_isList 
                              (x4::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_72 :
   forall x y, 
    (DP_R_xml_0_non_scc_72 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_19;
        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_73  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U41(tt,V2_)),isNeList(V2_)> *)
    | DP_R_xml_0_non_scc_73_0 :
     forall x4 x12, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Term 
         algebra.F.id_tt nil)::x4::nil)) x12) ->
       DP_R_xml_0_non_scc_73 (algebra.Alg.Term algebra.F.id_isNeList 
                              (x4::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_73 :
   forall x y, 
    (DP_R_xml_0_non_scc_73 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_74  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U41(tt,V2_)),U42(isNeList(V2_))> *)
    | DP_R_xml_0_non_scc_74_0 :
     forall x4 x12, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Term 
         algebra.F.id_tt nil)::x4::nil)) x12) ->
       DP_R_xml_0_non_scc_74 (algebra.Alg.Term algebra.F.id_U42 
                              ((algebra.Alg.Term algebra.F.id_isNeList 
                              (x4::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_74 :
   forall x y, 
    (DP_R_xml_0_non_scc_74 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_23;
        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_75  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U21(tt,V2_)),isList(V2_)> *)
    | DP_R_xml_0_non_scc_75_0 :
     forall x4 x12, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Term 
         algebra.F.id_tt nil)::x4::nil)) x12) ->
       DP_R_xml_0_non_scc_75 (algebra.Alg.Term algebra.F.id_isList (x4::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_75 :
   forall x y, 
    (DP_R_xml_0_non_scc_75 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_76  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(U21(tt,V2_)),U22(isList(V2_))> *)
    | DP_R_xml_0_non_scc_76_0 :
     forall x4 x12, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Term 
         algebra.F.id_tt nil)::x4::nil)) x12) ->
       DP_R_xml_0_non_scc_76 (algebra.Alg.Term algebra.F.id_U22 
                              ((algebra.Alg.Term algebra.F.id_isList 
                              (x4::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_76 :
   forall x y, 
    (DP_R_xml_0_non_scc_76 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_29;
        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_77  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(__(__(X_,Y_),Z_)),__(Y_,Z_)> *)
    | DP_R_xml_0_non_scc_77_0 :
     forall x12 x2 x1 x3, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Term algebra.F.id___ 
         (x1::x2::nil))::x3::nil)) x12) ->
       DP_R_xml_0_non_scc_77 (algebra.Alg.Term algebra.F.id___ (x2::x3::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_77 :
   forall x y, 
    (DP_R_xml_0_non_scc_77 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_35;
        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_78  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(__(__(X_,Y_),Z_)),__(X_,__(Y_,Z_))> *)
    | DP_R_xml_0_non_scc_78_0 :
     forall x12 x2 x1 x3, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Term algebra.F.id___ 
         (x1::x2::nil))::x3::nil)) x12) ->
       DP_R_xml_0_non_scc_78 (algebra.Alg.Term algebra.F.id___ (x1::
                              (algebra.Alg.Term algebra.F.id___ (x2::
                              x3::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_78 :
   forall x y, 
    (DP_R_xml_0_non_scc_78 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_35;
        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_79  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <active(__(X1_,X2_)),active(X2_)> *)
    | DP_R_xml_0_scc_79_0 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
       DP_R_xml_0_scc_79 (algebra.Alg.Term algebra.F.id_active (x10::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
     (* <active(__(X1_,X2_)),active(X1_)> *)
    | DP_R_xml_0_scc_79_1 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x12) ->
       DP_R_xml_0_scc_79 (algebra.Alg.Term algebra.F.id_active (x9::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
     (* <active(U11(X_)),active(X_)> *)
    | DP_R_xml_0_scc_79_2 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U11 (x1::nil)) x12) ->
       DP_R_xml_0_scc_79 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
     (* <active(U21(X1_,X2_)),active(X1_)> *)
    | DP_R_xml_0_scc_79_3 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U21 (x9::x10::nil)) x12) ->
       DP_R_xml_0_scc_79 (algebra.Alg.Term algebra.F.id_active (x9::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
     (* <active(U22(X_)),active(X_)> *)
    | DP_R_xml_0_scc_79_4 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U22 (x1::nil)) x12) ->
       DP_R_xml_0_scc_79 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
     (* <active(U31(X_)),active(X_)> *)
    | DP_R_xml_0_scc_79_5 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U31 (x1::nil)) x12) ->
       DP_R_xml_0_scc_79 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
     (* <active(U41(X1_,X2_)),active(X1_)> *)
    | DP_R_xml_0_scc_79_6 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U41 (x9::x10::nil)) x12) ->
       DP_R_xml_0_scc_79 (algebra.Alg.Term algebra.F.id_active (x9::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
     (* <active(U42(X_)),active(X_)> *)
    | DP_R_xml_0_scc_79_7 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U42 (x1::nil)) x12) ->
       DP_R_xml_0_scc_79 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
     (* <active(U51(X1_,X2_)),active(X1_)> *)
    | DP_R_xml_0_scc_79_8 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U51 (x9::x10::nil)) x12) ->
       DP_R_xml_0_scc_79 (algebra.Alg.Term algebra.F.id_active (x9::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
     (* <active(U52(X_)),active(X_)> *)
    | DP_R_xml_0_scc_79_9 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U52 (x1::nil)) x12) ->
       DP_R_xml_0_scc_79 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
     (* <active(U61(X_)),active(X_)> *)
    | DP_R_xml_0_scc_79_10 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U61 (x1::nil)) x12) ->
       DP_R_xml_0_scc_79 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
     (* <active(U71(X1_,X2_)),active(X1_)> *)
    | DP_R_xml_0_scc_79_11 :
     forall x12 x10 x9, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x12) ->
       DP_R_xml_0_scc_79 (algebra.Alg.Term algebra.F.id_active (x9::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
     (* <active(U72(X_)),active(X_)> *)
    | DP_R_xml_0_scc_79_12 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) x12) ->
       DP_R_xml_0_scc_79 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
     (* <active(U81(X_)),active(X_)> *)
    | DP_R_xml_0_scc_79_13 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_U81 (x1::nil)) x12) ->
       DP_R_xml_0_scc_79 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_active (x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_79.
   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 (x12:Z) := 3* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 1 + 2* x12.
   
   Definition P_id_isList (x12:Z) := 3 + 1* x12.
   
   Definition P_id_i  := 0.
   
   Definition P_id_U11 (x12:Z) := 1 + 2* x12.
   
   Definition P_id_isQid (x12:Z) := 2.
   
   Definition P_id_isNeList (x12:Z) := 3.
   
   Definition P_id_ok (x12:Z) := 0.
   
   Definition P_id_mark (x12:Z) := 0.
   
   Definition P_id_isPal (x12:Z) := 0.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 3 + 1* x12.
   
   Definition P_id_u  := 0.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 1 + 2* x12.
   
   Definition P_id_a  := 0.
   
   Definition P_id_U52 (x12:Z) := 1 + 2* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 3 + 2* x12 + 1* x13.
   
   Definition P_id_U72 (x12:Z) := 3 + 2* x12.
   
   Definition P_id_U31 (x12:Z) := 3 + 1* x12.
   
   Definition P_id_o  := 0.
   
   Definition P_id_tt  := 2.
   
   Definition P_id_isNePal (x12:Z) := 0.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 3 + 1* x12.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_U81 (x12:Z) := 1 + 1* x12.
   
   Definition P_id_U42 (x12:Z) := 3 + 2* x12.
   
   Definition P_id_proper (x12:Z) := 1 + 2* x12.
   
   Definition P_id_U22 (x12:Z) := 3 + 2* x12.
   
   Definition P_id_e  := 1.
   
   Definition P_id_U61 (x12:Z) := 3 + 1* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 0.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 2* x12.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_79.
   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_79.
  
  Definition wf_DP_R_xml_0_scc_79  := WF_DP_R_xml_0_scc_79.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_79 :
   forall x y, (DP_R_xml_0_scc_79 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_79).
    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_78;
         econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
       ((eapply acc_DP_R_xml_0_non_scc_77;
          econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
        ((eapply acc_DP_R_xml_0_non_scc_76;
           econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
         ((eapply acc_DP_R_xml_0_non_scc_75;
            econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
          ((eapply acc_DP_R_xml_0_non_scc_74;
             econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
           ((eapply acc_DP_R_xml_0_non_scc_73;
              econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
            ((eapply acc_DP_R_xml_0_non_scc_72;
               econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
             ((eapply acc_DP_R_xml_0_non_scc_71;
                econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
              ((eapply acc_DP_R_xml_0_non_scc_70;
                 econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
               ((eapply acc_DP_R_xml_0_non_scc_69;
                  econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                ((eapply acc_DP_R_xml_0_non_scc_68;
                   econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                 ((eapply acc_DP_R_xml_0_non_scc_67;
                    econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                  ((eapply acc_DP_R_xml_0_non_scc_66;
                     econstructor 
                     (eassumption)||(algebra.Alg_ext.star_refl' ))||
                   ((eapply acc_DP_R_xml_0_non_scc_65;
                      econstructor 
                      (eassumption)||(algebra.Alg_ext.star_refl' ))||
                    ((eapply acc_DP_R_xml_0_non_scc_64;
                       econstructor 
                       (eassumption)||(algebra.Alg_ext.star_refl' ))||
                     ((eapply acc_DP_R_xml_0_non_scc_63;
                        econstructor 
                        (eassumption)||(algebra.Alg_ext.star_refl' ))||
                      ((eapply acc_DP_R_xml_0_non_scc_62;
                         econstructor 
                         (eassumption)||(algebra.Alg_ext.star_refl' ))||
                       ((eapply acc_DP_R_xml_0_non_scc_61;
                          econstructor 
                          (eassumption)||(algebra.Alg_ext.star_refl' ))||
                        ((eapply acc_DP_R_xml_0_non_scc_60;
                           econstructor 
                           (eassumption)||(algebra.Alg_ext.star_refl' ))||
                         ((eapply acc_DP_R_xml_0_non_scc_59;
                            econstructor 
                            (eassumption)||(algebra.Alg_ext.star_refl' ))||
                          ((eapply acc_DP_R_xml_0_non_scc_58;
                             econstructor 
                             (eassumption)||(algebra.Alg_ext.star_refl' ))||
                           ((eapply acc_DP_R_xml_0_non_scc_57;
                              econstructor 
                              (eassumption)||(algebra.Alg_ext.star_refl' ))||
                            ((eapply acc_DP_R_xml_0_non_scc_56;
                               econstructor 
                               (eassumption)||(algebra.Alg_ext.star_refl' ))||
                             ((eapply acc_DP_R_xml_0_non_scc_55;
                                econstructor 
                                (eassumption)||(algebra.Alg_ext.star_refl' ))||
                              ((eapply acc_DP_R_xml_0_non_scc_54;
                                 econstructor 
                                 (eassumption)||(algebra.Alg_ext.star_refl' ))||
                               ((eapply acc_DP_R_xml_0_non_scc_53;
                                  econstructor 
                                  (eassumption)||
                                  (algebra.Alg_ext.star_refl' ))||
                                ((eapply acc_DP_R_xml_0_non_scc_52;
                                   econstructor 
                                   (eassumption)||
                                   (algebra.Alg_ext.star_refl' ))||
                                 ((eapply acc_DP_R_xml_0_non_scc_51;
                                    econstructor 
                                    (eassumption)||
                                    (algebra.Alg_ext.star_refl' ))||
                                  ((eapply acc_DP_R_xml_0_non_scc_50;
                                     econstructor 
                                     (eassumption)||
                                     (algebra.Alg_ext.star_refl' ))||
                                   ((eapply acc_DP_R_xml_0_non_scc_49;
                                      econstructor 
                                      (eassumption)||
                                      (algebra.Alg_ext.star_refl' ))||
                                    ((eapply acc_DP_R_xml_0_non_scc_48;
                                       econstructor 
                                       (eassumption)||
                                       (algebra.Alg_ext.star_refl' ))||
                                     ((eapply acc_DP_R_xml_0_non_scc_47;
                                        econstructor 
                                        (eassumption)||
                                        (algebra.Alg_ext.star_refl' ))||
                                      ((eapply acc_DP_R_xml_0_non_scc_46;
                                         econstructor 
                                         (eassumption)||
                                         (algebra.Alg_ext.star_refl' ))||
                                       ((eapply acc_DP_R_xml_0_non_scc_45;
                                          econstructor 
                                          (eassumption)||
                                          (algebra.Alg_ext.star_refl' ))||
                                        ((eapply acc_DP_R_xml_0_non_scc_44;
                                           econstructor 
                                           (eassumption)||
                                           (algebra.Alg_ext.star_refl' ))||
                                         ((eapply acc_DP_R_xml_0_non_scc_43;
                                            econstructor 
                                            (eassumption)||
                                            (algebra.Alg_ext.star_refl' ))||
                                          ((eapply acc_DP_R_xml_0_non_scc_42;
                                             econstructor 
                                             (eassumption)||
                                             (algebra.Alg_ext.star_refl' ))||
                                           ((eapply acc_DP_R_xml_0_non_scc_41;
                                              econstructor 
                                              (eassumption)||
                                              (algebra.Alg_ext.star_refl' ))||
                                            ((eapply acc_DP_R_xml_0_non_scc_40;
                                               
                                               econstructor 
                                               (eassumption)||
                                               (algebra.Alg_ext.star_refl' ))||
                                             ((eapply acc_DP_R_xml_0_non_scc_39;
                                                
                                                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_79.
  Qed.
  
  
  Inductive DP_R_xml_0_non_scc_80  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <top(ok(X_)),active(X_)> *)
    | DP_R_xml_0_non_scc_80_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
       DP_R_xml_0_non_scc_80 (algebra.Alg.Term algebra.F.id_active (x1::nil)) 
        (algebra.Alg.Term algebra.F.id_top (x12::nil))
  .
  
  
  Lemma acc_DP_R_xml_0_non_scc_80 :
   forall x y, 
    (DP_R_xml_0_non_scc_80 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_79;
        econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
      ((eapply acc_DP_R_xml_0_non_scc_78;
         econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
       ((eapply acc_DP_R_xml_0_non_scc_77;
          econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
        ((eapply acc_DP_R_xml_0_non_scc_76;
           econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
         ((eapply acc_DP_R_xml_0_non_scc_75;
            econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
          ((eapply acc_DP_R_xml_0_non_scc_74;
             econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
           ((eapply acc_DP_R_xml_0_non_scc_73;
              econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
            ((eapply acc_DP_R_xml_0_non_scc_72;
               econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
             ((eapply acc_DP_R_xml_0_non_scc_71;
                econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
              ((eapply acc_DP_R_xml_0_non_scc_70;
                 econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
               ((eapply acc_DP_R_xml_0_non_scc_69;
                  econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                ((eapply acc_DP_R_xml_0_non_scc_68;
                   econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                 ((eapply acc_DP_R_xml_0_non_scc_67;
                    econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                  ((eapply acc_DP_R_xml_0_non_scc_66;
                     econstructor 
                     (eassumption)||(algebra.Alg_ext.star_refl' ))||
                   ((eapply acc_DP_R_xml_0_non_scc_65;
                      econstructor 
                      (eassumption)||(algebra.Alg_ext.star_refl' ))||
                    ((eapply acc_DP_R_xml_0_non_scc_64;
                       econstructor 
                       (eassumption)||(algebra.Alg_ext.star_refl' ))||
                     ((eapply acc_DP_R_xml_0_non_scc_63;
                        econstructor 
                        (eassumption)||(algebra.Alg_ext.star_refl' ))||
                      ((eapply acc_DP_R_xml_0_non_scc_62;
                         econstructor 
                         (eassumption)||(algebra.Alg_ext.star_refl' ))||
                       ((eapply acc_DP_R_xml_0_non_scc_61;
                          econstructor 
                          (eassumption)||(algebra.Alg_ext.star_refl' ))||
                        ((eapply acc_DP_R_xml_0_non_scc_60;
                           econstructor 
                           (eassumption)||(algebra.Alg_ext.star_refl' ))||
                         ((eapply acc_DP_R_xml_0_non_scc_59;
                            econstructor 
                            (eassumption)||(algebra.Alg_ext.star_refl' ))||
                          ((eapply acc_DP_R_xml_0_non_scc_58;
                             econstructor 
                             (eassumption)||(algebra.Alg_ext.star_refl' ))||
                           ((eapply acc_DP_R_xml_0_non_scc_57;
                              econstructor 
                              (eassumption)||(algebra.Alg_ext.star_refl' ))||
                            ((eapply acc_DP_R_xml_0_non_scc_56;
                               econstructor 
                               (eassumption)||(algebra.Alg_ext.star_refl' ))||
                             ((eapply acc_DP_R_xml_0_non_scc_55;
                                econstructor 
                                (eassumption)||(algebra.Alg_ext.star_refl' ))||
                              ((eapply acc_DP_R_xml_0_non_scc_54;
                                 econstructor 
                                 (eassumption)||(algebra.Alg_ext.star_refl' ))||
                               ((eapply acc_DP_R_xml_0_non_scc_53;
                                  econstructor 
                                  (eassumption)||
                                  (algebra.Alg_ext.star_refl' ))||
                                ((eapply acc_DP_R_xml_0_non_scc_52;
                                   econstructor 
                                   (eassumption)||
                                   (algebra.Alg_ext.star_refl' ))||
                                 ((eapply acc_DP_R_xml_0_non_scc_51;
                                    econstructor 
                                    (eassumption)||
                                    (algebra.Alg_ext.star_refl' ))||
                                  ((eapply acc_DP_R_xml_0_non_scc_50;
                                     econstructor 
                                     (eassumption)||
                                     (algebra.Alg_ext.star_refl' ))||
                                   ((eapply acc_DP_R_xml_0_non_scc_49;
                                      econstructor 
                                      (eassumption)||
                                      (algebra.Alg_ext.star_refl' ))||
                                    ((eapply acc_DP_R_xml_0_non_scc_48;
                                       econstructor 
                                       (eassumption)||
                                       (algebra.Alg_ext.star_refl' ))||
                                     ((eapply acc_DP_R_xml_0_non_scc_47;
                                        econstructor 
                                        (eassumption)||
                                        (algebra.Alg_ext.star_refl' ))||
                                      ((eapply acc_DP_R_xml_0_non_scc_46;
                                         econstructor 
                                         (eassumption)||
                                         (algebra.Alg_ext.star_refl' ))||
                                       ((eapply acc_DP_R_xml_0_non_scc_45;
                                          econstructor 
                                          (eassumption)||
                                          (algebra.Alg_ext.star_refl' ))||
                                        ((eapply acc_DP_R_xml_0_non_scc_44;
                                           econstructor 
                                           (eassumption)||
                                           (algebra.Alg_ext.star_refl' ))||
                                         ((eapply acc_DP_R_xml_0_non_scc_43;
                                            econstructor 
                                            (eassumption)||
                                            (algebra.Alg_ext.star_refl' ))||
                                          ((eapply acc_DP_R_xml_0_non_scc_42;
                                             econstructor 
                                             (eassumption)||
                                             (algebra.Alg_ext.star_refl' ))||
                                           ((eapply acc_DP_R_xml_0_non_scc_41;
                                              econstructor 
                                              (eassumption)||
                                              (algebra.Alg_ext.star_refl' ))||
                                            ((eapply acc_DP_R_xml_0_non_scc_40;
                                               
                                               econstructor 
                                               (eassumption)||
                                               (algebra.Alg_ext.star_refl' ))||
                                             ((eapply acc_DP_R_xml_0_non_scc_39;
                                                
                                                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_81  :
   algebra.Alg.term ->algebra.Alg.term ->Prop := 
     (* <top(ok(X_)),top(active(X_))> *)
    | DP_R_xml_0_scc_81_0 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
       DP_R_xml_0_scc_81 (algebra.Alg.Term algebra.F.id_top 
                          ((algebra.Alg.Term algebra.F.id_active 
                          (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_top (x12::nil))
     (* <top(mark(X_)),top(proper(X_))> *)
    | DP_R_xml_0_scc_81_1 :
     forall x12 x1, 
      (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                 
        (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
       DP_R_xml_0_scc_81 (algebra.Alg.Term algebra.F.id_top 
                          ((algebra.Alg.Term algebra.F.id_proper 
                          (x1::nil))::nil)) 
        (algebra.Alg.Term algebra.F.id_top (x12::nil))
  .
  
  
  Module WF_DP_R_xml_0_scc_81.
   Inductive DP_R_xml_0_scc_81_large  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <top(ok(X_)),top(active(X_))> *)
     | DP_R_xml_0_scc_81_large_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_ok (x1::nil)) x12) ->
        DP_R_xml_0_scc_81_large (algebra.Alg.Term algebra.F.id_top 
                                 ((algebra.Alg.Term algebra.F.id_active 
                                 (x1::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_top (x12::nil))
   .
   
   
   Inductive DP_R_xml_0_scc_81_strict  :
    algebra.Alg.term ->algebra.Alg.term ->Prop := 
      (* <top(mark(X_)),top(proper(X_))> *)
     | DP_R_xml_0_scc_81_strict_0 :
      forall x12 x1, 
       (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules)
                                  
         (algebra.Alg.Term algebra.F.id_mark (x1::nil)) x12) ->
        DP_R_xml_0_scc_81_strict (algebra.Alg.Term algebra.F.id_top 
                                  ((algebra.Alg.Term algebra.F.id_proper 
                                  (x1::nil))::nil)) 
         (algebra.Alg.Term algebra.F.id_top (x12::nil))
   .
   
   
   Module WF_DP_R_xml_0_scc_81_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 (x12:Z) := 1* x12.
    
    Definition P_id_U71 (x12:Z) (x13:Z) := 3* x13.
    
    Definition P_id_isList (x12:Z) := 2 + 3* x12.
    
    Definition P_id_i  := 2.
    
    Definition P_id_U11 (x12:Z) := 2 + 2* x12.
    
    Definition P_id_isQid (x12:Z) := 1 + 1* x12.
    
    Definition P_id_isNeList (x12:Z) := 1* x12.
    
    Definition P_id_ok (x12:Z) := 1 + 1* x12.
    
    Definition P_id_mark (x12:Z) := 0.
    
    Definition P_id_isPal (x12:Z) := 2* x12.
    
    Definition P_id_U41 (x12:Z) (x13:Z) := 1 + 2* x12 + 3* x13.
    
    Definition P_id_u  := 2.
    
    Definition P_id_U21 (x12:Z) (x13:Z) := 2 + 3* x12 + 3* x13.
    
    Definition P_id_a  := 2.
    
    Definition P_id_U52 (x12:Z) := 2* x12.
    
    Definition P_id___ (x12:Z) (x13:Z) := 1* x12.
    
    Definition P_id_U72 (x12:Z) := 2* x12.
    
    Definition P_id_U31 (x12:Z) := 2* x12.
    
    Definition P_id_o  := 2.
    
    Definition P_id_tt  := 2.
    
    Definition P_id_isNePal (x12:Z) := 2* x12.
    
    Definition P_id_U51 (x12:Z) (x13:Z) := 3 + 3* x13.
    
    Definition P_id_top (x12:Z) := 0.
    
    Definition P_id_nil  := 2.
    
    Definition P_id_U81 (x12:Z) := 1* x12.
    
    Definition P_id_U42 (x12:Z) := 1* x12.
    
    Definition P_id_proper (x12:Z) := 2* x12.
    
    Definition P_id_U22 (x12:Z) := 2* x12.
    
    Definition P_id_e  := 2.
    
    Definition P_id_U61 (x12:Z) := 2* x12.
    
    Lemma P_id_active_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_isList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
    Proof.
      intros x13 x12.
      intros [H_1 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 x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U72_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_top_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
    Proof.
      intros x13 x12.
      intros [H_1 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_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_active x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_i_bounded : 0 <= P_id_i .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNeList_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
    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 x12, (0 <= x12) ->0 <= P_id_mark x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_u_bounded : 0 <= P_id_u .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_a_bounded : 0 <= P_id_a .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_o_bounded : 0 <= P_id_o .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_tt_bounded : 0 <= P_id_tt .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_isNePal_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_bounded :
     forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
    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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_proper_bounded :
     forall x12, (0 <= x12) ->0 <= P_id_proper x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_e_bounded : 0 <= P_id_e .
    Proof.
      intros .
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
    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_U71 P_id_isList P_id_i P_id_U11 
       P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
       P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
       P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
       P_id_proper P_id_U22 P_id_e P_id_U61.
    
    Lemma measure_equation :
     forall t, 
      measure t = match t with
                    | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                     P_id_active (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                     P_id_U71 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                     P_id_isList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                    | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                     P_id_U11 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                     P_id_isQid (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                     P_id_isNeList (measure x12)
                    | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                     P_id_ok (measure x12)
                    | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                     P_id_mark (measure x12)
                    | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                     P_id_isPal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                     P_id_U41 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                    | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                     P_id_U21 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                    | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                     P_id_U52 (measure x12)
                    | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                     P_id___ (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                     P_id_U72 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                     P_id_U31 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                    | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                    | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                     P_id_isNePal (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                     P_id_U51 (measure x13) (measure x12)
                    | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                     P_id_top (measure x12)
                    | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                    | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                     P_id_U81 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                     P_id_U42 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                     P_id_proper (measure x12)
                    | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                     P_id_U22 (measure x12)
                    | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                    | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                     P_id_U61 (measure x12)
                    | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
    Qed.
    
    Definition P_id_U22_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNEPAL (x12:Z) := 0.
    
    Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U52_hat_1 (x12:Z) := 0.
    
    Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U42_hat_1 (x12:Z) := 0.
    
    Definition P_id_TOP (x12:Z) := 1* x12.
    
    Definition P_id_ISQID (x12:Z) := 0.
    
    Definition P_id_ISPAL (x12:Z) := 0.
    
    Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_ACTIVE (x12:Z) := 0.
    
    Definition P_id_ISLIST (x12:Z) := 0.
    
    Definition P_id_PROPER (x12:Z) := 0.
    
    Definition P_id_U31_hat_1 (x12:Z) := 0.
    
    Definition P_id_U72_hat_1 (x12:Z) := 0.
    
    Definition P_id_U61_hat_1 (x12:Z) := 0.
    
    Definition P_id_ISNELIST (x12:Z) := 0.
    
    Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
    
    Definition P_id_U11_hat_1 (x12:Z) := 0.
    
    Definition P_id_U81_hat_1 (x12:Z) := 0.
    
    Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
    
    Lemma P_id_U22_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNEPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U21_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U52_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U51_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U42_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_TOP_monotonic :
     forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISQID_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISPAL_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U71_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISLIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_PROPER_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U31_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U72_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U61_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_ISNELIST_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U41_hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      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_U11_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id_U81_hat_1_monotonic :
     forall x12 x13, 
      (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
    Proof.
      intros x13 x12.
      intros [H_1 H_0].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Lemma P_id____hat_1_monotonic :
     forall x12 x14 x13 x15, 
      (0 <= x15)/\ (x15 <= x14) ->
       (0 <= x13)/\ (x13 <= x12) ->
        P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
    Proof.
      intros x15 x14 x13 x12.
      intros [H_1 H_0].
      intros [H_3 H_2].
      
      cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
       (auto with zarith)||(repeat (translate_vars );prove_ineq ).
    Qed.
    
    Definition marked_measure  := 
      InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
       P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
       P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
       P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
       P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
       P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
       P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
       P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
       P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
       P_id_U81_hat_1 P_id____hat_1.
    
    Lemma marked_measure_equation :
     forall t, 
      marked_measure t = match t with
                           | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                            P_id_U22_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNePal 
                              (x12::nil)) =>
                            P_id_ISNEPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U21 (x13::
                              x12::nil)) =>
                            P_id_U21_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                            P_id_U52_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U51 (x13::
                              x12::nil)) =>
                            P_id_U51_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                            P_id_U42_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                            P_id_TOP (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                            P_id_ISQID (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                            P_id_ISPAL (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U71 (x13::
                              x12::nil)) =>
                            P_id_U71_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_active 
                              (x12::nil)) =>
                            P_id_ACTIVE (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isList 
                              (x12::nil)) =>
                            P_id_ISLIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_proper 
                              (x12::nil)) =>
                            P_id_PROPER (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                            P_id_U31_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                            P_id_U72_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                            P_id_U61_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_isNeList 
                              (x12::nil)) =>
                            P_id_ISNELIST (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U41 (x13::
                              x12::nil)) =>
                            P_id_U41_hat_1 (measure x13) (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                            P_id_U11_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                            P_id_U81_hat_1 (measure x12)
                           | (algebra.Alg.Term algebra.F.id___ (x13::
                              x12::nil)) =>
                            P_id____hat_1 (measure x13) (measure x12)
                           | _ => 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_U71_monotonic;assumption.
      intros ;apply P_id_isList_monotonic;assumption.
      intros ;apply P_id_U11_monotonic;assumption.
      intros ;apply P_id_isQid_monotonic;assumption.
      intros ;apply P_id_isNeList_monotonic;assumption.
      intros ;apply P_id_ok_monotonic;assumption.
      intros ;apply P_id_mark_monotonic;assumption.
      intros ;apply P_id_isPal_monotonic;assumption.
      intros ;apply P_id_U41_monotonic;assumption.
      intros ;apply P_id_U21_monotonic;assumption.
      intros ;apply P_id_U52_monotonic;assumption.
      intros ;apply P_id____monotonic;assumption.
      intros ;apply P_id_U72_monotonic;assumption.
      intros ;apply P_id_U31_monotonic;assumption.
      intros ;apply P_id_isNePal_monotonic;assumption.
      intros ;apply P_id_U51_monotonic;assumption.
      intros ;apply P_id_top_monotonic;assumption.
      intros ;apply P_id_U81_monotonic;assumption.
      intros ;apply P_id_U42_monotonic;assumption.
      intros ;apply P_id_proper_monotonic;assumption.
      intros ;apply P_id_U22_monotonic;assumption.
      intros ;apply P_id_U61_monotonic;assumption.
      intros ;apply P_id_active_bounded;assumption.
      intros ;apply P_id_U71_bounded;assumption.
      intros ;apply P_id_isList_bounded;assumption.
      intros ;apply P_id_i_bounded;assumption.
      intros ;apply P_id_U11_bounded;assumption.
      intros ;apply P_id_isQid_bounded;assumption.
      intros ;apply P_id_isNeList_bounded;assumption.
      intros ;apply P_id_ok_bounded;assumption.
      intros ;apply P_id_mark_bounded;assumption.
      intros ;apply P_id_isPal_bounded;assumption.
      intros ;apply P_id_U41_bounded;assumption.
      intros ;apply P_id_u_bounded;assumption.
      intros ;apply P_id_U21_bounded;assumption.
      intros ;apply P_id_a_bounded;assumption.
      intros ;apply P_id_U52_bounded;assumption.
      intros ;apply P_id____bounded;assumption.
      intros ;apply P_id_U72_bounded;assumption.
      intros ;apply P_id_U31_bounded;assumption.
      intros ;apply P_id_o_bounded;assumption.
      intros ;apply P_id_tt_bounded;assumption.
      intros ;apply P_id_isNePal_bounded;assumption.
      intros ;apply P_id_U51_bounded;assumption.
      intros ;apply P_id_top_bounded;assumption.
      intros ;apply P_id_nil_bounded;assumption.
      intros ;apply P_id_U81_bounded;assumption.
      intros ;apply P_id_U42_bounded;assumption.
      intros ;apply P_id_proper_bounded;assumption.
      intros ;apply P_id_U22_bounded;assumption.
      intros ;apply P_id_e_bounded;assumption.
      intros ;apply P_id_U61_bounded;assumption.
      apply rules_monotonic.
      intros ;apply P_id_U22_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNEPAL_monotonic;assumption.
      intros ;apply P_id_U21_hat_1_monotonic;assumption.
      intros ;apply P_id_U52_hat_1_monotonic;assumption.
      intros ;apply P_id_U51_hat_1_monotonic;assumption.
      intros ;apply P_id_U42_hat_1_monotonic;assumption.
      intros ;apply P_id_TOP_monotonic;assumption.
      intros ;apply P_id_ISQID_monotonic;assumption.
      intros ;apply P_id_ISPAL_monotonic;assumption.
      intros ;apply P_id_U71_hat_1_monotonic;assumption.
      intros ;apply P_id_ACTIVE_monotonic;assumption.
      intros ;apply P_id_ISLIST_monotonic;assumption.
      intros ;apply P_id_PROPER_monotonic;assumption.
      intros ;apply P_id_U31_hat_1_monotonic;assumption.
      intros ;apply P_id_U72_hat_1_monotonic;assumption.
      intros ;apply P_id_U61_hat_1_monotonic;assumption.
      intros ;apply P_id_ISNELIST_monotonic;assumption.
      intros ;apply P_id_U41_hat_1_monotonic;assumption.
      intros ;apply P_id_U11_hat_1_monotonic;assumption.
      intros ;apply P_id_U81_hat_1_monotonic;assumption.
      intros ;apply P_id____hat_1_monotonic;assumption.
    Qed.
    
    Ltac rewrite_and_unfold  :=
     do 2 (rewrite marked_measure_equation);
      repeat (
      match goal with
        |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
         rewrite (measure_equation (algebra.Alg.Term f t))
        | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
         rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
        end
      ).
    
    
    Lemma wf : well_founded WF_DP_R_xml_0_scc_81.DP_R_xml_0_scc_81_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_81_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 (x12:Z) := 1* x12.
   
   Definition P_id_U71 (x12:Z) (x13:Z) := 2 + 3* x12 + 2* x13.
   
   Definition P_id_isList (x12:Z) := 3 + 3* x12.
   
   Definition P_id_i  := 3.
   
   Definition P_id_U11 (x12:Z) := 1 + 1* x12.
   
   Definition P_id_isQid (x12:Z) := 1* x12.
   
   Definition P_id_isNeList (x12:Z) := 1 + 3* x12.
   
   Definition P_id_ok (x12:Z) := 1* x12.
   
   Definition P_id_mark (x12:Z) := 1 + 1* x12.
   
   Definition P_id_isPal (x12:Z) := 3 + 2* x12.
   
   Definition P_id_U41 (x12:Z) (x13:Z) := 3 + 1* x12 + 3* x13.
   
   Definition P_id_u  := 3.
   
   Definition P_id_U21 (x12:Z) (x13:Z) := 1 + 2* x12 + 3* x13.
   
   Definition P_id_a  := 3.
   
   Definition P_id_U52 (x12:Z) := 2 + 1* x12.
   
   Definition P_id___ (x12:Z) (x13:Z) := 2 + 3* x12 + 1* x13.
   
   Definition P_id_U72 (x12:Z) := 2 + 1* x12.
   
   Definition P_id_U31 (x12:Z) := 2* x12.
   
   Definition P_id_o  := 3.
   
   Definition P_id_tt  := 2.
   
   Definition P_id_isNePal (x12:Z) := 1 + 2* x12.
   
   Definition P_id_U51 (x12:Z) (x13:Z) := 2 + 2* x12 + 3* x13.
   
   Definition P_id_top (x12:Z) := 0.
   
   Definition P_id_nil  := 0.
   
   Definition P_id_U81 (x12:Z) := 1 + 1* x12.
   
   Definition P_id_U42 (x12:Z) := 2 + 1* x12.
   
   Definition P_id_proper (x12:Z) := 1* x12.
   
   Definition P_id_U22 (x12:Z) := 1 + 1* x12.
   
   Definition P_id_e  := 3.
   
   Definition P_id_U61 (x12:Z) := 2* x12.
   
   Lemma P_id_active_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_active x13 <= P_id_active x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U71 x13 x15 <= P_id_U71 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_isList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isList x13 <= P_id_isList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U11 x13 <= P_id_U11 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isQid x13 <= P_id_isQid x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNeList x13 <= P_id_isNeList x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_ok x13 <= P_id_ok x12.
   Proof.
     intros x13 x12.
     intros [H_1 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 x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_mark x13 <= P_id_mark x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isPal x13 <= P_id_isPal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U41 x13 x15 <= P_id_U41 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U21 x13 x15 <= P_id_U21 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U52 x13 <= P_id_U52 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id___ x13 x15 <= P_id___ x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U72_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U72 x13 <= P_id_U72 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U31 x13 <= P_id_U31 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_isNePal x13 <= P_id_isNePal x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->P_id_U51 x13 x15 <= P_id_U51 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_top_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_top x13 <= P_id_top x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U81 x13 <= P_id_U81 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U42 x13 <= P_id_U42 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_proper x13 <= P_id_proper x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U22 x13 <= P_id_U22 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_U61 x13 <= P_id_U61 x12.
   Proof.
     intros x13 x12.
     intros [H_1 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_bounded : forall x12, (0 <= x12) ->0 <= P_id_active x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U71 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isList_bounded : forall x12, (0 <= x12) ->0 <= P_id_isList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_i_bounded : 0 <= P_id_i .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U11_bounded : forall x12, (0 <= x12) ->0 <= P_id_U11 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isQid_bounded : forall x12, (0 <= x12) ->0 <= P_id_isQid x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNeList_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNeList x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ok_bounded : forall x12, (0 <= x12) ->0 <= P_id_ok x12.
   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 x12, (0 <= x12) ->0 <= P_id_mark x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isPal_bounded : forall x12, (0 <= x12) ->0 <= P_id_isPal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_u_bounded : 0 <= P_id_u .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U21 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_a_bounded : 0 <= P_id_a .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U52_bounded : forall x12, (0 <= x12) ->0 <= P_id_U52 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id___ x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_bounded : forall x12, (0 <= x12) ->0 <= P_id_U72 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_bounded : forall x12, (0 <= x12) ->0 <= P_id_U31 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_o_bounded : 0 <= P_id_o .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_tt_bounded : 0 <= P_id_tt .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_isNePal_bounded :
    forall x12, (0 <= x12) ->0 <= P_id_isNePal x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_bounded :
    forall x12 x13, (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_top_bounded : forall x12, (0 <= x12) ->0 <= P_id_top x12.
   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_U81_bounded : forall x12, (0 <= x12) ->0 <= P_id_U81 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_bounded : forall x12, (0 <= x12) ->0 <= P_id_U42 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_proper_bounded : forall x12, (0 <= x12) ->0 <= P_id_proper x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U22_bounded : forall x12, (0 <= x12) ->0 <= P_id_U22 x12.
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_e_bounded : 0 <= P_id_e .
   Proof.
     intros .
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_bounded : forall x12, (0 <= x12) ->0 <= P_id_U61 x12.
   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_U71 P_id_isList P_id_i P_id_U11 
      P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal P_id_U41 
      P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 P_id_o 
      P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 P_id_U42 
      P_id_proper P_id_U22 P_id_e P_id_U61.
   
   Lemma measure_equation :
    forall t, 
     measure t = match t with
                   | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                    P_id_active (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U71 (x13::x12::nil)) =>
                    P_id_U71 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                    P_id_isList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i 
                   | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                    P_id_U11 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                    P_id_isQid (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isNeList (x12::nil)) =>
                    P_id_isNeList (measure x12)
                   | (algebra.Alg.Term algebra.F.id_ok (x12::nil)) =>
                    P_id_ok (measure x12)
                   | (algebra.Alg.Term algebra.F.id_mark (x12::nil)) =>
                    P_id_mark (measure x12)
                   | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                    P_id_isPal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::nil)) =>
                    P_id_U41 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u 
                   | (algebra.Alg.Term algebra.F.id_U21 (x13::x12::nil)) =>
                    P_id_U21 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a 
                   | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                    P_id_U52 (measure x12)
                   | (algebra.Alg.Term algebra.F.id___ (x13::x12::nil)) =>
                    P_id___ (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                    P_id_U72 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                    P_id_U31 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o 
                   | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt 
                   | (algebra.Alg.Term algebra.F.id_isNePal (x12::nil)) =>
                    P_id_isNePal (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::nil)) =>
                    P_id_U51 (measure x13) (measure x12)
                   | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                    P_id_top (measure x12)
                   | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil 
                   | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                    P_id_U81 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                    P_id_U42 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                    P_id_proper (measure x12)
                   | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                    P_id_U22 (measure x12)
                   | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e 
                   | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                    P_id_U61 (measure x12)
                   | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
   Qed.
   
   Definition P_id_U22_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNEPAL (x12:Z) := 0.
   
   Definition P_id_U21_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U52_hat_1 (x12:Z) := 0.
   
   Definition P_id_U51_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U42_hat_1 (x12:Z) := 0.
   
   Definition P_id_TOP (x12:Z) := 1* x12.
   
   Definition P_id_ISQID (x12:Z) := 0.
   
   Definition P_id_ISPAL (x12:Z) := 0.
   
   Definition P_id_U71_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_ACTIVE (x12:Z) := 0.
   
   Definition P_id_ISLIST (x12:Z) := 0.
   
   Definition P_id_PROPER (x12:Z) := 0.
   
   Definition P_id_U31_hat_1 (x12:Z) := 0.
   
   Definition P_id_U72_hat_1 (x12:Z) := 0.
   
   Definition P_id_U61_hat_1 (x12:Z) := 0.
   
   Definition P_id_ISNELIST (x12:Z) := 0.
   
   Definition P_id_U41_hat_1 (x12:Z) (x13:Z) := 0.
   
   Definition P_id_U11_hat_1 (x12:Z) := 0.
   
   Definition P_id_U81_hat_1 (x12:Z) := 0.
   
   Definition P_id____hat_1 (x12:Z) (x13:Z) := 0.
   
   Lemma P_id_U22_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U22_hat_1 x13 <= P_id_U22_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNEPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNEPAL x13 <= P_id_ISNEPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U21_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U21_hat_1 x13 x15 <= P_id_U21_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U52_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U52_hat_1 x13 <= P_id_U52_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U51_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U51_hat_1 x13 x15 <= P_id_U51_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U42_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U42_hat_1 x13 <= P_id_U42_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_TOP_monotonic :
    forall x12 x13, (0 <= x13)/\ (x13 <= x12) ->P_id_TOP x13 <= P_id_TOP x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISQID_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISQID x13 <= P_id_ISQID x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISPAL_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISPAL x13 <= P_id_ISPAL x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U71_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U71_hat_1 x13 x15 <= P_id_U71_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ACTIVE x13 <= P_id_ACTIVE x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISLIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISLIST x13 <= P_id_ISLIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_PROPER_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_PROPER x13 <= P_id_PROPER x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U31_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U31_hat_1 x13 <= P_id_U31_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U72_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U72_hat_1 x13 <= P_id_U72_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U61_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U61_hat_1 x13 <= P_id_U61_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_ISNELIST_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_ISNELIST x13 <= P_id_ISNELIST x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U41_hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id_U41_hat_1 x13 x15 <= P_id_U41_hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     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_U11_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U11_hat_1 x13 <= P_id_U11_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id_U81_hat_1_monotonic :
    forall x12 x13, 
     (0 <= x13)/\ (x13 <= x12) ->P_id_U81_hat_1 x13 <= P_id_U81_hat_1 x12.
   Proof.
     intros x13 x12.
     intros [H_1 H_0].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Lemma P_id____hat_1_monotonic :
    forall x12 x14 x13 x15, 
     (0 <= x15)/\ (x15 <= x14) ->
      (0 <= x13)/\ (x13 <= x12) ->
       P_id____hat_1 x13 x15 <= P_id____hat_1 x12 x14.
   Proof.
     intros x15 x14 x13 x12.
     intros [H_1 H_0].
     intros [H_3 H_2].
     
     cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition;
      (auto with zarith)||(repeat (translate_vars );prove_ineq ).
   Qed.
   
   Definition marked_measure  := 
     InterpZ.marked_measure 0 P_id_active P_id_U71 P_id_isList P_id_i 
      P_id_U11 P_id_isQid P_id_isNeList P_id_ok P_id_mark P_id_isPal 
      P_id_U41 P_id_u P_id_U21 P_id_a P_id_U52 P_id___ P_id_U72 P_id_U31 
      P_id_o P_id_tt P_id_isNePal P_id_U51 P_id_top P_id_nil P_id_U81 
      P_id_U42 P_id_proper P_id_U22 P_id_e P_id_U61 P_id_U22_hat_1 
      P_id_ISNEPAL P_id_U21_hat_1 P_id_U52_hat_1 P_id_U51_hat_1 
      P_id_U42_hat_1 P_id_TOP P_id_ISQID P_id_ISPAL P_id_U71_hat_1 
      P_id_ACTIVE P_id_ISLIST P_id_PROPER P_id_U31_hat_1 P_id_U72_hat_1 
      P_id_U61_hat_1 P_id_ISNELIST P_id_U41_hat_1 P_id_U11_hat_1 
      P_id_U81_hat_1 P_id____hat_1.
   
   Lemma marked_measure_equation :
    forall t, 
     marked_measure t = match t with
                          | (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) =>
                           P_id_U22_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNePal 
                             (x12::nil)) =>
                           P_id_ISNEPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U21 (x13::
                             x12::nil)) =>
                           P_id_U21_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U52 (x12::nil)) =>
                           P_id_U52_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U51 (x13::
                             x12::nil)) =>
                           P_id_U51_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U42 (x12::nil)) =>
                           P_id_U42_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_top (x12::nil)) =>
                           P_id_TOP (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isQid (x12::nil)) =>
                           P_id_ISQID (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isPal (x12::nil)) =>
                           P_id_ISPAL (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U71 (x13::
                             x12::nil)) =>
                           P_id_U71_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_active (x12::nil)) =>
                           P_id_ACTIVE (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isList (x12::nil)) =>
                           P_id_ISLIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_proper (x12::nil)) =>
                           P_id_PROPER (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U31 (x12::nil)) =>
                           P_id_U31_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U72 (x12::nil)) =>
                           P_id_U72_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U61 (x12::nil)) =>
                           P_id_U61_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_isNeList 
                             (x12::nil)) =>
                           P_id_ISNELIST (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U41 (x13::
                             x12::nil)) =>
                           P_id_U41_hat_1 (measure x13) (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U11 (x12::nil)) =>
                           P_id_U11_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id_U81 (x12::nil)) =>
                           P_id_U81_hat_1 (measure x12)
                          | (algebra.Alg.Term algebra.F.id___ (x13::
                             x12::nil)) =>
                           P_id____hat_1 (measure x13) (measure x12)
                          | _ => 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_U71_monotonic;assumption.
     intros ;apply P_id_isList_monotonic;assumption.
     intros ;apply P_id_U11_monotonic;assumption.
     intros ;apply P_id_isQid_monotonic;assumption.
     intros ;apply P_id_isNeList_monotonic;assumption.
     intros ;apply P_id_ok_monotonic;assumption.
     intros ;apply P_id_mark_monotonic;assumption.
     intros ;apply P_id_isPal_monotonic;assumption.
     intros ;apply P_id_U41_monotonic;assumption.
     intros ;apply P_id_U21_monotonic;assumption.
     intros ;apply P_id_U52_monotonic;assumption.
     intros ;apply P_id____monotonic;assumption.
     intros ;apply P_id_U72_monotonic;assumption.
     intros ;apply P_id_U31_monotonic;assumption.
     intros ;apply P_id_isNePal_monotonic;assumption.
     intros ;apply P_id_U51_monotonic;assumption.
     intros ;apply P_id_top_monotonic;assumption.
     intros ;apply P_id_U81_monotonic;assumption.
     intros ;apply P_id_U42_monotonic;assumption.
     intros ;apply P_id_proper_monotonic;assumption.
     intros ;apply P_id_U22_monotonic;assumption.
     intros ;apply P_id_U61_monotonic;assumption.
     intros ;apply P_id_active_bounded;assumption.
     intros ;apply P_id_U71_bounded;assumption.
     intros ;apply P_id_isList_bounded;assumption.
     intros ;apply P_id_i_bounded;assumption.
     intros ;apply P_id_U11_bounded;assumption.
     intros ;apply P_id_isQid_bounded;assumption.
     intros ;apply P_id_isNeList_bounded;assumption.
     intros ;apply P_id_ok_bounded;assumption.
     intros ;apply P_id_mark_bounded;assumption.
     intros ;apply P_id_isPal_bounded;assumption.
     intros ;apply P_id_U41_bounded;assumption.
     intros ;apply P_id_u_bounded;assumption.
     intros ;apply P_id_U21_bounded;assumption.
     intros ;apply P_id_a_bounded;assumption.
     intros ;apply P_id_U52_bounded;assumption.
     intros ;apply P_id____bounded;assumption.
     intros ;apply P_id_U72_bounded;assumption.
     intros ;apply P_id_U31_bounded;assumption.
     intros ;apply P_id_o_bounded;assumption.
     intros ;apply P_id_tt_bounded;assumption.
     intros ;apply P_id_isNePal_bounded;assumption.
     intros ;apply P_id_U51_bounded;assumption.
     intros ;apply P_id_top_bounded;assumption.
     intros ;apply P_id_nil_bounded;assumption.
     intros ;apply P_id_U81_bounded;assumption.
     intros ;apply P_id_U42_bounded;assumption.
     intros ;apply P_id_proper_bounded;assumption.
     intros ;apply P_id_U22_bounded;assumption.
     intros ;apply P_id_e_bounded;assumption.
     intros ;apply P_id_U61_bounded;assumption.
     apply rules_monotonic.
     intros ;apply P_id_U22_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNEPAL_monotonic;assumption.
     intros ;apply P_id_U21_hat_1_monotonic;assumption.
     intros ;apply P_id_U52_hat_1_monotonic;assumption.
     intros ;apply P_id_U51_hat_1_monotonic;assumption.
     intros ;apply P_id_U42_hat_1_monotonic;assumption.
     intros ;apply P_id_TOP_monotonic;assumption.
     intros ;apply P_id_ISQID_monotonic;assumption.
     intros ;apply P_id_ISPAL_monotonic;assumption.
     intros ;apply P_id_U71_hat_1_monotonic;assumption.
     intros ;apply P_id_ACTIVE_monotonic;assumption.
     intros ;apply P_id_ISLIST_monotonic;assumption.
     intros ;apply P_id_PROPER_monotonic;assumption.
     intros ;apply P_id_U31_hat_1_monotonic;assumption.
     intros ;apply P_id_U72_hat_1_monotonic;assumption.
     intros ;apply P_id_U61_hat_1_monotonic;assumption.
     intros ;apply P_id_ISNELIST_monotonic;assumption.
     intros ;apply P_id_U41_hat_1_monotonic;assumption.
     intros ;apply P_id_U11_hat_1_monotonic;assumption.
     intros ;apply P_id_U81_hat_1_monotonic;assumption.
     intros ;apply P_id____hat_1_monotonic;assumption.
   Qed.
   
   Ltac rewrite_and_unfold  :=
    do 2 (rewrite marked_measure_equation);
     repeat (
     match goal with
       |  |- context [measure (algebra.Alg.Term ?f ?t)] =>
        rewrite (measure_equation (algebra.Alg.Term f t))
       | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ =>
        rewrite (measure_equation (algebra.Alg.Term f t)) in H|-
       end
     ).
   
   Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b).
   
   Definition le a b := marked_measure a <= marked_measure b.
   
   Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c.
   Proof.
     unfold lt, le in *.
     intros a b c.
     apply (interp.le_lt_compat_right (interp.o_Z 0)).
   Qed.
   
   Lemma wf_lt : well_founded lt.
   Proof.
     unfold lt in *.
     apply Inverse_Image.wf_inverse_image with  (B:=Z).
     apply Zwf.Zwf_well_founded.
   Qed.
   
   Lemma DP_R_xml_0_scc_81_strict_in_lt :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_81_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_81_large_in_le :
    Relation_Definitions.inclusion _ DP_R_xml_0_scc_81_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_81_large  := WF_DP_R_xml_0_scc_81_large.wf.
   
   
   Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_81.
   Proof.
     intros x.
     apply (well_founded_ind wf_lt).
     clear x.
     intros x.
     pattern x.
     apply (@Acc_ind _ DP_R_xml_0_scc_81_large).
     clear x.
     intros x _ IHx IHx'.
     constructor.
     intros y H.
     
     destruct H;
      (apply IHx';apply DP_R_xml_0_scc_81_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_81_large_in_le;econstructor eassumption])).
     apply wf_DP_R_xml_0_scc_81_large.
   Qed.
  End WF_DP_R_xml_0_scc_81.
  
  Definition wf_DP_R_xml_0_scc_81  := WF_DP_R_xml_0_scc_81.wf.
  
  
  Lemma acc_DP_R_xml_0_scc_81 :
   forall x y, (DP_R_xml_0_scc_81 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_81).
    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_80;
         econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
       ((eapply acc_DP_R_xml_0_non_scc_38;
          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_81.
  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_80;
       econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
     ((eapply acc_DP_R_xml_0_non_scc_79;
        econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
      ((eapply acc_DP_R_xml_0_non_scc_78;
         econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
       ((eapply acc_DP_R_xml_0_non_scc_77;
          econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
        ((eapply acc_DP_R_xml_0_non_scc_76;
           econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
         ((eapply acc_DP_R_xml_0_non_scc_75;
            econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
          ((eapply acc_DP_R_xml_0_non_scc_74;
             econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
           ((eapply acc_DP_R_xml_0_non_scc_73;
              econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
            ((eapply acc_DP_R_xml_0_non_scc_72;
               econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
             ((eapply acc_DP_R_xml_0_non_scc_71;
                econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
              ((eapply acc_DP_R_xml_0_non_scc_70;
                 econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
               ((eapply acc_DP_R_xml_0_non_scc_69;
                  econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                ((eapply acc_DP_R_xml_0_non_scc_68;
                   econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                 ((eapply acc_DP_R_xml_0_non_scc_67;
                    econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))||
                  ((eapply acc_DP_R_xml_0_non_scc_66;
                     econstructor 
                     (eassumption)||(algebra.Alg_ext.star_refl' ))||
                   ((eapply acc_DP_R_xml_0_non_scc_65;
                      econstructor 
                      (eassumption)||(algebra.Alg_ext.star_refl' ))||
                    ((eapply acc_DP_R_xml_0_non_scc_64;
                       econstructor 
                       (eassumption)||(algebra.Alg_ext.star_refl' ))||
                     ((eapply acc_DP_R_xml_0_non_scc_63;
                        econstructor 
                        (eassumption)||(algebra.Alg_ext.star_refl' ))||
                      ((eapply acc_DP_R_xml_0_non_scc_62;
                         econstructor 
                         (eassumption)||(algebra.Alg_ext.star_refl' ))||
                       ((eapply acc_DP_R_xml_0_non_scc_61;
                          econstructor 
                          (eassumption)||(algebra.Alg_ext.star_refl' ))||
                        ((eapply acc_DP_R_xml_0_non_scc_60;
                           econstructor 
                           (eassumption)||(algebra.Alg_ext.star_refl' ))||
                         ((eapply acc_DP_R_xml_0_non_scc_59;
                            econstructor 
                            (eassumption)||(algebra.Alg_ext.star_refl' ))||
                          ((eapply acc_DP_R_xml_0_non_scc_58;
                             econstructor 
                             (eassumption)||(algebra.Alg_ext.star_refl' ))||
                           ((eapply acc_DP_R_xml_0_non_scc_57;
                              econstructor 
                              (eassumption)||(algebra.Alg_ext.star_refl' ))||
                            ((eapply acc_DP_R_xml_0_non_scc_56;
                               econstructor 
                               (eassumption)||(algebra.Alg_ext.star_refl' ))||
                             ((eapply acc_DP_R_xml_0_non_scc_55;
                                econstructor 
                                (eassumption)||(algebra.Alg_ext.star_refl' ))||
                              ((eapply acc_DP_R_xml_0_non_scc_54;
                                 econstructor 
                                 (eassumption)||(algebra.Alg_ext.star_refl' ))||
                               ((eapply acc_DP_R_xml_0_non_scc_53;
                                  econstructor 
                                  (eassumption)||
                                  (algebra.Alg_ext.star_refl' ))||
                                ((eapply acc_DP_R_xml_0_non_scc_52;
                                   econstructor 
                                   (eassumption)||
                                   (algebra.Alg_ext.star_refl' ))||
                                 ((eapply acc_DP_R_xml_0_non_scc_51;
                                    econstructor 
                                    (eassumption)||
                                    (algebra.Alg_ext.star_refl' ))||
                                  ((eapply acc_DP_R_xml_0_non_scc_50;
                                     econstructor 
                                     (eassumption)||
                                     (algebra.Alg_ext.star_refl' ))||
                                   ((eapply acc_DP_R_xml_0_non_scc_49;
                                      econstructor 
                                      (eassumption)||
                                      (algebra.Alg_ext.star_refl' ))||
                                    ((eapply acc_DP_R_xml_0_non_scc_48;
                                       econstructor 
                                       (eassumption)||
                                       (algebra.Alg_ext.star_refl' ))||
                                     ((eapply acc_DP_R_xml_0_non_scc_47;
                                        econstructor 
                                        (eassumption)||
                                        (algebra.Alg_ext.star_refl' ))||
                                      ((eapply acc_DP_R_xml_0_non_scc_46;
                                         econstructor 
                                         (eassumption)||
                                         (algebra.Alg_ext.star_refl' ))||
                                       ((eapply acc_DP_R_xml_0_non_scc_45;
                                          econstructor 
                                          (eassumption)||
                                          (algebra.Alg_ext.star_refl' ))||
                                        ((eapply acc_DP_R_xml_0_non_scc_44;
                                           econstructor 
                                           (eassumption)||
                                           (algebra.Alg_ext.star_refl' ))||
                                         ((eapply acc_DP_R_xml_0_non_scc_43;
                                            econstructor 
                                            (eassumption)||
                                            (algebra.Alg_ext.star_refl' ))||
                                          ((eapply acc_DP_R_xml_0_non_scc_42;
                                             econstructor 
                                             (eassumption)||
                                             (algebra.Alg_ext.star_refl' ))||
                                           ((eapply acc_DP_R_xml_0_non_scc_41;
                                              econstructor 
                                              (eassumption)||
                                              (algebra.Alg_ext.star_refl' ))||
                                            ((eapply acc_DP_R_xml_0_non_scc_40;
                                               
                                               econstructor 
                                               (eassumption)||
                                               (algebra.Alg_ext.star_refl' ))||
                                             ((eapply acc_DP_R_xml_0_non_scc_39;
                                                
                                                econstructor 
                                                (eassumption)||
                                                (algebra.Alg_ext.star_refl' ))||
                                              ((eapply acc_DP_R_xml_0_non_scc_38;
                                                 
                                                 econstructor 
                                                 (eassumption)||
                                                 (algebra.Alg_ext.star_refl' ))||
                                               ((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_81;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_80;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_79;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_78;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_77;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_76;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_75;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_74;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_73;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_72;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_71;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_70;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_69;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_68;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_67;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_66;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_65;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_64;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_63;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_62;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_61;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_60;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_59;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_58;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_57;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_56;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_55;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_54;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_53;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_52;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_51;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_50;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_49;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_48;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_47;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_46;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_45;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_44;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_43;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_42;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_41;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_40;
                                                                    
                                                                    econstructor 
                                                                    (
                                                                    eassumption)||
                                                                    (
                                                                    algebra.Alg_ext.star_refl' ))||
                                                                    (
                                                                    (
                                                                    eapply 
                                                                    acc_DP_R_xml_0_scc_39;
                                                                    
                                                                    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___PALINDROME_nokinds_noand_C.trs/a3pat.v" ***
*** End: ***
 *)