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_a__zeros *) | id_a__zeros : symb (* id_isNatKind *) | id_isNatKind : symb (* id_a__isNatIList *) | id_a__isNatIList : symb (* id_a__U21 *) | id_a__U21 : symb (* id_U41 *) | id_U41 : symb (* id_a__and *) | id_a__and : symb (* id_a__U11 *) | id_a__U11 : symb (* id_U21 *) | id_U21 : symb (* id_a__U61 *) | id_a__U61 : symb (* id_a__U32 *) | id_a__U32 : symb (* id_U51 *) | id_U51 : symb (* id_isNatIListKind *) | id_isNatIListKind : symb (* id_0 *) | id_0 : symb (* id_U12 *) | id_U12 : symb (* id_a__U52 *) | id_a__U52 : symb (* id_a__isNat *) | id_a__isNat : symb (* id_U43 *) | id_U43 : symb (* id_a__isNatIListKind *) | id_a__isNatIListKind : symb (* id_a__U12 *) | id_a__U12 : symb (* id_U31 *) | id_U31 : symb (* id_a__length *) | id_a__length : symb (* id_a__U42 *) | id_a__U42 : symb (* id_U53 *) | id_U53 : symb (* id_and *) | id_and : symb (* id_cons *) | id_cons : symb (* id_U11 *) | id_U11 : symb (* id_a__U51 *) | id_a__U51 : symb (* id_a__U22 *) | id_a__U22 : symb (* id_U42 *) | id_U42 : symb (* id_length *) | id_length : symb (* id_tt *) | id_tt : symb (* id_U22 *) | id_U22 : symb (* id_s *) | id_s : symb (* id_a__U41 *) | id_a__U41 : symb (* id_U52 *) | id_U52 : symb (* id_nil *) | id_nil : symb (* id_zeros *) | id_zeros : symb (* id_isNatList *) | id_isNatList : symb (* id_a__U53 *) | id_a__U53 : symb (* id_a__U31 *) | id_a__U31 : symb (* id_isNatIList *) | id_isNatIList : symb (* id_a__isNatKind *) | id_a__isNatKind : symb (* id_a__isNatList *) | id_a__isNatList : symb (* id_U32 *) | id_U32 : symb (* id_mark *) | id_mark : symb (* id_a__U43 *) | id_a__U43 : symb (* id_U61 *) | id_U61 : symb (* id_isNat *) | id_isNat : symb . Definition symb_eq_bool (f1 f2:symb) : bool := match f1,f2 with | id_a__zeros,id_a__zeros => true | id_isNatKind,id_isNatKind => true | id_a__isNatIList,id_a__isNatIList => true | id_a__U21,id_a__U21 => true | id_U41,id_U41 => true | id_a__and,id_a__and => true | id_a__U11,id_a__U11 => true | id_U21,id_U21 => true | id_a__U61,id_a__U61 => true | id_a__U32,id_a__U32 => true | id_U51,id_U51 => true | id_isNatIListKind,id_isNatIListKind => true | id_0,id_0 => true | id_U12,id_U12 => true | id_a__U52,id_a__U52 => true | id_a__isNat,id_a__isNat => true | id_U43,id_U43 => true | id_a__isNatIListKind,id_a__isNatIListKind => true | id_a__U12,id_a__U12 => true | id_U31,id_U31 => true | id_a__length,id_a__length => true | id_a__U42,id_a__U42 => true | id_U53,id_U53 => true | id_and,id_and => true | id_cons,id_cons => true | id_U11,id_U11 => true | id_a__U51,id_a__U51 => true | id_a__U22,id_a__U22 => true | id_U42,id_U42 => true | id_length,id_length => true | id_tt,id_tt => true | id_U22,id_U22 => true | id_s,id_s => true | id_a__U41,id_a__U41 => true | id_U52,id_U52 => true | id_nil,id_nil => true | id_zeros,id_zeros => true | id_isNatList,id_isNatList => true | id_a__U53,id_a__U53 => true | id_a__U31,id_a__U31 => true | id_isNatIList,id_isNatIList => true | id_a__isNatKind,id_a__isNatKind => true | id_a__isNatList,id_a__isNatList => true | id_U32,id_U32 => true | id_mark,id_mark => true | id_a__U43,id_a__U43 => true | id_U61,id_U61 => true | id_isNat,id_isNat => 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_a__zeros,id_a__zeros => refl_equal _ | id_isNatKind,id_isNatKind => refl_equal _ | id_a__isNatIList,id_a__isNatIList => refl_equal _ | id_a__U21,id_a__U21 => refl_equal _ | id_U41,id_U41 => refl_equal _ | id_a__and,id_a__and => refl_equal _ | id_a__U11,id_a__U11 => refl_equal _ | id_U21,id_U21 => refl_equal _ | id_a__U61,id_a__U61 => refl_equal _ | id_a__U32,id_a__U32 => refl_equal _ | id_U51,id_U51 => refl_equal _ | id_isNatIListKind,id_isNatIListKind => refl_equal _ | id_0,id_0 => refl_equal _ | id_U12,id_U12 => refl_equal _ | id_a__U52,id_a__U52 => refl_equal _ | id_a__isNat,id_a__isNat => refl_equal _ | id_U43,id_U43 => refl_equal _ | id_a__isNatIListKind,id_a__isNatIListKind => refl_equal _ | id_a__U12,id_a__U12 => refl_equal _ | id_U31,id_U31 => refl_equal _ | id_a__length,id_a__length => refl_equal _ | id_a__U42,id_a__U42 => refl_equal _ | id_U53,id_U53 => refl_equal _ | id_and,id_and => refl_equal _ | id_cons,id_cons => refl_equal _ | id_U11,id_U11 => refl_equal _ | id_a__U51,id_a__U51 => refl_equal _ | id_a__U22,id_a__U22 => refl_equal _ | id_U42,id_U42 => refl_equal _ | id_length,id_length => refl_equal _ | id_tt,id_tt => refl_equal _ | id_U22,id_U22 => refl_equal _ | id_s,id_s => refl_equal _ | id_a__U41,id_a__U41 => refl_equal _ | id_U52,id_U52 => refl_equal _ | id_nil,id_nil => refl_equal _ | id_zeros,id_zeros => refl_equal _ | id_isNatList,id_isNatList => refl_equal _ | id_a__U53,id_a__U53 => refl_equal _ | id_a__U31,id_a__U31 => refl_equal _ | id_isNatIList,id_isNatIList => refl_equal _ | id_a__isNatKind,id_a__isNatKind => refl_equal _ | id_a__isNatList,id_a__isNatList => refl_equal _ | id_U32,id_U32 => refl_equal _ | id_mark,id_mark => refl_equal _ | id_a__U43,id_a__U43 => refl_equal _ | id_U61,id_U61 => refl_equal _ | id_isNat,id_isNat => refl_equal _ | _,_ => _ end;intros abs;discriminate. Defined. Definition arity (f:symb) := match f with | id_a__zeros => term.Free 0 | id_isNatKind => term.Free 1 | id_a__isNatIList => term.Free 1 | id_a__U21 => term.Free 2 | id_U41 => term.Free 3 | id_a__and => term.Free 2 | id_a__U11 => term.Free 2 | id_U21 => term.Free 2 | id_a__U61 => term.Free 2 | id_a__U32 => term.Free 1 | id_U51 => term.Free 3 | id_isNatIListKind => term.Free 1 | id_0 => term.Free 0 | id_U12 => term.Free 1 | id_a__U52 => term.Free 2 | id_a__isNat => term.Free 1 | id_U43 => term.Free 1 | id_a__isNatIListKind => term.Free 1 | id_a__U12 => term.Free 1 | id_U31 => term.Free 2 | id_a__length => term.Free 1 | id_a__U42 => term.Free 2 | id_U53 => term.Free 1 | id_and => term.Free 2 | id_cons => term.Free 2 | id_U11 => term.Free 2 | id_a__U51 => term.Free 3 | id_a__U22 => term.Free 1 | id_U42 => term.Free 2 | id_length => term.Free 1 | id_tt => term.Free 0 | id_U22 => term.Free 1 | id_s => term.Free 1 | id_a__U41 => term.Free 3 | id_U52 => term.Free 2 | id_nil => term.Free 0 | id_zeros => term.Free 0 | id_isNatList => term.Free 1 | id_a__U53 => term.Free 1 | id_a__U31 => term.Free 2 | id_isNatIList => term.Free 1 | id_a__isNatKind => term.Free 1 | id_a__isNatList => term.Free 1 | id_U32 => term.Free 1 | id_mark => term.Free 1 | id_a__U43 => term.Free 1 | id_U61 => term.Free 2 | id_isNat => term.Free 1 end. Definition symb_order (f1 f2:symb) : bool := match f1,f2 with | id_a__zeros,id_a__zeros => true | id_a__zeros,id_isNatKind => false | id_a__zeros,id_a__isNatIList => false | id_a__zeros,id_a__U21 => false | id_a__zeros,id_U41 => false | id_a__zeros,id_a__and => false | id_a__zeros,id_a__U11 => false | id_a__zeros,id_U21 => false | id_a__zeros,id_a__U61 => false | id_a__zeros,id_a__U32 => false | id_a__zeros,id_U51 => false | id_a__zeros,id_isNatIListKind => false | id_a__zeros,id_0 => false | id_a__zeros,id_U12 => false | id_a__zeros,id_a__U52 => false | id_a__zeros,id_a__isNat => false | id_a__zeros,id_U43 => false | id_a__zeros,id_a__isNatIListKind => false | id_a__zeros,id_a__U12 => false | id_a__zeros,id_U31 => false | id_a__zeros,id_a__length => false | id_a__zeros,id_a__U42 => false | id_a__zeros,id_U53 => false | id_a__zeros,id_and => false | id_a__zeros,id_cons => false | id_a__zeros,id_U11 => false | id_a__zeros,id_a__U51 => false | id_a__zeros,id_a__U22 => false | id_a__zeros,id_U42 => false | id_a__zeros,id_length => false | id_a__zeros,id_tt => false | id_a__zeros,id_U22 => false | id_a__zeros,id_s => false | id_a__zeros,id_a__U41 => false | id_a__zeros,id_U52 => false | id_a__zeros,id_nil => false | id_a__zeros,id_zeros => false | id_a__zeros,id_isNatList => false | id_a__zeros,id_a__U53 => false | id_a__zeros,id_a__U31 => false | id_a__zeros,id_isNatIList => false | id_a__zeros,id_a__isNatKind => false | id_a__zeros,id_a__isNatList => false | id_a__zeros,id_U32 => false | id_a__zeros,id_mark => false | id_a__zeros,id_a__U43 => false | id_a__zeros,id_U61 => false | id_a__zeros,id_isNat => false | id_isNatKind,id_a__zeros => true | id_isNatKind,id_isNatKind => true | id_isNatKind,id_a__isNatIList => false | id_isNatKind,id_a__U21 => false | id_isNatKind,id_U41 => false | id_isNatKind,id_a__and => false | id_isNatKind,id_a__U11 => false | id_isNatKind,id_U21 => false | id_isNatKind,id_a__U61 => false | id_isNatKind,id_a__U32 => false | id_isNatKind,id_U51 => false | id_isNatKind,id_isNatIListKind => false | id_isNatKind,id_0 => false | id_isNatKind,id_U12 => false | id_isNatKind,id_a__U52 => false | id_isNatKind,id_a__isNat => false | id_isNatKind,id_U43 => false | id_isNatKind,id_a__isNatIListKind => false | id_isNatKind,id_a__U12 => false | id_isNatKind,id_U31 => false | id_isNatKind,id_a__length => false | id_isNatKind,id_a__U42 => false | id_isNatKind,id_U53 => false | id_isNatKind,id_and => false | id_isNatKind,id_cons => false | id_isNatKind,id_U11 => false | id_isNatKind,id_a__U51 => false | id_isNatKind,id_a__U22 => false | id_isNatKind,id_U42 => false | id_isNatKind,id_length => false | id_isNatKind,id_tt => false | id_isNatKind,id_U22 => false | id_isNatKind,id_s => false | id_isNatKind,id_a__U41 => false | id_isNatKind,id_U52 => false | id_isNatKind,id_nil => false | id_isNatKind,id_zeros => false | id_isNatKind,id_isNatList => false | id_isNatKind,id_a__U53 => false | id_isNatKind,id_a__U31 => false | id_isNatKind,id_isNatIList => false | id_isNatKind,id_a__isNatKind => false | id_isNatKind,id_a__isNatList => false | id_isNatKind,id_U32 => false | id_isNatKind,id_mark => false | id_isNatKind,id_a__U43 => false | id_isNatKind,id_U61 => false | id_isNatKind,id_isNat => false | id_a__isNatIList,id_a__zeros => true | id_a__isNatIList,id_isNatKind => true | id_a__isNatIList,id_a__isNatIList => true | id_a__isNatIList,id_a__U21 => false | id_a__isNatIList,id_U41 => false | id_a__isNatIList,id_a__and => false | id_a__isNatIList,id_a__U11 => false | id_a__isNatIList,id_U21 => false | id_a__isNatIList,id_a__U61 => false | id_a__isNatIList,id_a__U32 => false | id_a__isNatIList,id_U51 => false | id_a__isNatIList,id_isNatIListKind => false | id_a__isNatIList,id_0 => false | id_a__isNatIList,id_U12 => false | id_a__isNatIList,id_a__U52 => false | id_a__isNatIList,id_a__isNat => false | id_a__isNatIList,id_U43 => false | id_a__isNatIList,id_a__isNatIListKind => false | id_a__isNatIList,id_a__U12 => false | id_a__isNatIList,id_U31 => false | id_a__isNatIList,id_a__length => false | id_a__isNatIList,id_a__U42 => false | id_a__isNatIList,id_U53 => false | id_a__isNatIList,id_and => false | id_a__isNatIList,id_cons => false | id_a__isNatIList,id_U11 => false | id_a__isNatIList,id_a__U51 => false | id_a__isNatIList,id_a__U22 => false | id_a__isNatIList,id_U42 => false | id_a__isNatIList,id_length => false | id_a__isNatIList,id_tt => false | id_a__isNatIList,id_U22 => false | id_a__isNatIList,id_s => false | id_a__isNatIList,id_a__U41 => false | id_a__isNatIList,id_U52 => false | id_a__isNatIList,id_nil => false | id_a__isNatIList,id_zeros => false | id_a__isNatIList,id_isNatList => false | id_a__isNatIList,id_a__U53 => false | id_a__isNatIList,id_a__U31 => false | id_a__isNatIList,id_isNatIList => false | id_a__isNatIList,id_a__isNatKind => false | id_a__isNatIList,id_a__isNatList => false | id_a__isNatIList,id_U32 => false | id_a__isNatIList,id_mark => false | id_a__isNatIList,id_a__U43 => false | id_a__isNatIList,id_U61 => false | id_a__isNatIList,id_isNat => false | id_a__U21,id_a__zeros => true | id_a__U21,id_isNatKind => true | id_a__U21,id_a__isNatIList => true | id_a__U21,id_a__U21 => true | id_a__U21,id_U41 => false | id_a__U21,id_a__and => false | id_a__U21,id_a__U11 => false | id_a__U21,id_U21 => false | id_a__U21,id_a__U61 => false | id_a__U21,id_a__U32 => false | id_a__U21,id_U51 => false | id_a__U21,id_isNatIListKind => false | id_a__U21,id_0 => false | id_a__U21,id_U12 => false | id_a__U21,id_a__U52 => false | id_a__U21,id_a__isNat => false | id_a__U21,id_U43 => false | id_a__U21,id_a__isNatIListKind => false | id_a__U21,id_a__U12 => false | id_a__U21,id_U31 => false | id_a__U21,id_a__length => false | id_a__U21,id_a__U42 => false | id_a__U21,id_U53 => false | id_a__U21,id_and => false | id_a__U21,id_cons => false | id_a__U21,id_U11 => false | id_a__U21,id_a__U51 => false | id_a__U21,id_a__U22 => false | id_a__U21,id_U42 => false | id_a__U21,id_length => false | id_a__U21,id_tt => false | id_a__U21,id_U22 => false | id_a__U21,id_s => false | id_a__U21,id_a__U41 => false | id_a__U21,id_U52 => false | id_a__U21,id_nil => false | id_a__U21,id_zeros => false | id_a__U21,id_isNatList => false | id_a__U21,id_a__U53 => false | id_a__U21,id_a__U31 => false | id_a__U21,id_isNatIList => false | id_a__U21,id_a__isNatKind => false | id_a__U21,id_a__isNatList => false | id_a__U21,id_U32 => false | id_a__U21,id_mark => false | id_a__U21,id_a__U43 => false | id_a__U21,id_U61 => false | id_a__U21,id_isNat => false | id_U41,id_a__zeros => true | id_U41,id_isNatKind => true | id_U41,id_a__isNatIList => true | id_U41,id_a__U21 => true | id_U41,id_U41 => true | id_U41,id_a__and => false | id_U41,id_a__U11 => false | id_U41,id_U21 => false | id_U41,id_a__U61 => false | id_U41,id_a__U32 => false | id_U41,id_U51 => false | id_U41,id_isNatIListKind => false | id_U41,id_0 => false | id_U41,id_U12 => false | id_U41,id_a__U52 => false | id_U41,id_a__isNat => false | id_U41,id_U43 => false | id_U41,id_a__isNatIListKind => false | id_U41,id_a__U12 => false | id_U41,id_U31 => false | id_U41,id_a__length => false | id_U41,id_a__U42 => false | id_U41,id_U53 => false | id_U41,id_and => false | id_U41,id_cons => false | id_U41,id_U11 => false | id_U41,id_a__U51 => false | id_U41,id_a__U22 => false | id_U41,id_U42 => false | id_U41,id_length => false | id_U41,id_tt => false | id_U41,id_U22 => false | id_U41,id_s => false | id_U41,id_a__U41 => false | id_U41,id_U52 => false | id_U41,id_nil => false | id_U41,id_zeros => false | id_U41,id_isNatList => false | id_U41,id_a__U53 => false | id_U41,id_a__U31 => false | id_U41,id_isNatIList => false | id_U41,id_a__isNatKind => false | id_U41,id_a__isNatList => false | id_U41,id_U32 => false | id_U41,id_mark => false | id_U41,id_a__U43 => false | id_U41,id_U61 => false | id_U41,id_isNat => false | id_a__and,id_a__zeros => true | id_a__and,id_isNatKind => true | id_a__and,id_a__isNatIList => true | id_a__and,id_a__U21 => true | id_a__and,id_U41 => true | id_a__and,id_a__and => true | id_a__and,id_a__U11 => false | id_a__and,id_U21 => false | id_a__and,id_a__U61 => false | id_a__and,id_a__U32 => false | id_a__and,id_U51 => false | id_a__and,id_isNatIListKind => false | id_a__and,id_0 => false | id_a__and,id_U12 => false | id_a__and,id_a__U52 => false | id_a__and,id_a__isNat => false | id_a__and,id_U43 => false | id_a__and,id_a__isNatIListKind => false | id_a__and,id_a__U12 => false | id_a__and,id_U31 => false | id_a__and,id_a__length => false | id_a__and,id_a__U42 => false | id_a__and,id_U53 => false | id_a__and,id_and => false | id_a__and,id_cons => false | id_a__and,id_U11 => false | id_a__and,id_a__U51 => false | id_a__and,id_a__U22 => false | id_a__and,id_U42 => false | id_a__and,id_length => false | id_a__and,id_tt => false | id_a__and,id_U22 => false | id_a__and,id_s => false | id_a__and,id_a__U41 => false | id_a__and,id_U52 => false | id_a__and,id_nil => false | id_a__and,id_zeros => false | id_a__and,id_isNatList => false | id_a__and,id_a__U53 => false | id_a__and,id_a__U31 => false | id_a__and,id_isNatIList => false | id_a__and,id_a__isNatKind => false | id_a__and,id_a__isNatList => false | id_a__and,id_U32 => false | id_a__and,id_mark => false | id_a__and,id_a__U43 => false | id_a__and,id_U61 => false | id_a__and,id_isNat => false | id_a__U11,id_a__zeros => true | id_a__U11,id_isNatKind => true | id_a__U11,id_a__isNatIList => true | id_a__U11,id_a__U21 => true | id_a__U11,id_U41 => true | id_a__U11,id_a__and => true | id_a__U11,id_a__U11 => true | id_a__U11,id_U21 => false | id_a__U11,id_a__U61 => false | id_a__U11,id_a__U32 => false | id_a__U11,id_U51 => false | id_a__U11,id_isNatIListKind => false | id_a__U11,id_0 => false | id_a__U11,id_U12 => false | id_a__U11,id_a__U52 => false | id_a__U11,id_a__isNat => false | id_a__U11,id_U43 => false | id_a__U11,id_a__isNatIListKind => false | id_a__U11,id_a__U12 => false | id_a__U11,id_U31 => false | id_a__U11,id_a__length => false | id_a__U11,id_a__U42 => false | id_a__U11,id_U53 => false | id_a__U11,id_and => false | id_a__U11,id_cons => false | id_a__U11,id_U11 => false | id_a__U11,id_a__U51 => false | id_a__U11,id_a__U22 => false | id_a__U11,id_U42 => false | id_a__U11,id_length => false | id_a__U11,id_tt => false | id_a__U11,id_U22 => false | id_a__U11,id_s => false | id_a__U11,id_a__U41 => false | id_a__U11,id_U52 => false | id_a__U11,id_nil => false | id_a__U11,id_zeros => false | id_a__U11,id_isNatList => false | id_a__U11,id_a__U53 => false | id_a__U11,id_a__U31 => false | id_a__U11,id_isNatIList => false | id_a__U11,id_a__isNatKind => false | id_a__U11,id_a__isNatList => false | id_a__U11,id_U32 => false | id_a__U11,id_mark => false | id_a__U11,id_a__U43 => false | id_a__U11,id_U61 => false | id_a__U11,id_isNat => false | id_U21,id_a__zeros => true | id_U21,id_isNatKind => true | id_U21,id_a__isNatIList => true | id_U21,id_a__U21 => true | id_U21,id_U41 => true | id_U21,id_a__and => true | id_U21,id_a__U11 => true | id_U21,id_U21 => true | id_U21,id_a__U61 => false | id_U21,id_a__U32 => false | id_U21,id_U51 => false | id_U21,id_isNatIListKind => false | id_U21,id_0 => false | id_U21,id_U12 => false | id_U21,id_a__U52 => false | id_U21,id_a__isNat => false | id_U21,id_U43 => false | id_U21,id_a__isNatIListKind => false | id_U21,id_a__U12 => false | id_U21,id_U31 => false | id_U21,id_a__length => false | id_U21,id_a__U42 => false | id_U21,id_U53 => false | id_U21,id_and => false | id_U21,id_cons => false | id_U21,id_U11 => false | id_U21,id_a__U51 => false | id_U21,id_a__U22 => false | id_U21,id_U42 => false | id_U21,id_length => false | id_U21,id_tt => false | id_U21,id_U22 => false | id_U21,id_s => false | id_U21,id_a__U41 => false | id_U21,id_U52 => false | id_U21,id_nil => false | id_U21,id_zeros => false | id_U21,id_isNatList => false | id_U21,id_a__U53 => false | id_U21,id_a__U31 => false | id_U21,id_isNatIList => false | id_U21,id_a__isNatKind => false | id_U21,id_a__isNatList => false | id_U21,id_U32 => false | id_U21,id_mark => false | id_U21,id_a__U43 => false | id_U21,id_U61 => false | id_U21,id_isNat => false | id_a__U61,id_a__zeros => true | id_a__U61,id_isNatKind => true | id_a__U61,id_a__isNatIList => true | id_a__U61,id_a__U21 => true | id_a__U61,id_U41 => true | id_a__U61,id_a__and => true | id_a__U61,id_a__U11 => true | id_a__U61,id_U21 => true | id_a__U61,id_a__U61 => true | id_a__U61,id_a__U32 => false | id_a__U61,id_U51 => false | id_a__U61,id_isNatIListKind => false | id_a__U61,id_0 => false | id_a__U61,id_U12 => false | id_a__U61,id_a__U52 => false | id_a__U61,id_a__isNat => false | id_a__U61,id_U43 => false | id_a__U61,id_a__isNatIListKind => false | id_a__U61,id_a__U12 => false | id_a__U61,id_U31 => false | id_a__U61,id_a__length => false | id_a__U61,id_a__U42 => false | id_a__U61,id_U53 => false | id_a__U61,id_and => false | id_a__U61,id_cons => false | id_a__U61,id_U11 => false | id_a__U61,id_a__U51 => false | id_a__U61,id_a__U22 => false | id_a__U61,id_U42 => false | id_a__U61,id_length => false | id_a__U61,id_tt => false | id_a__U61,id_U22 => false | id_a__U61,id_s => false | id_a__U61,id_a__U41 => false | id_a__U61,id_U52 => false | id_a__U61,id_nil => false | id_a__U61,id_zeros => false | id_a__U61,id_isNatList => false | id_a__U61,id_a__U53 => false | id_a__U61,id_a__U31 => false | id_a__U61,id_isNatIList => false | id_a__U61,id_a__isNatKind => false | id_a__U61,id_a__isNatList => false | id_a__U61,id_U32 => false | id_a__U61,id_mark => false | id_a__U61,id_a__U43 => false | id_a__U61,id_U61 => false | id_a__U61,id_isNat => false | id_a__U32,id_a__zeros => true | id_a__U32,id_isNatKind => true | id_a__U32,id_a__isNatIList => true | id_a__U32,id_a__U21 => true | id_a__U32,id_U41 => true | id_a__U32,id_a__and => true | id_a__U32,id_a__U11 => true | id_a__U32,id_U21 => true | id_a__U32,id_a__U61 => true | id_a__U32,id_a__U32 => true | id_a__U32,id_U51 => false | id_a__U32,id_isNatIListKind => false | id_a__U32,id_0 => false | id_a__U32,id_U12 => false | id_a__U32,id_a__U52 => false | id_a__U32,id_a__isNat => false | id_a__U32,id_U43 => false | id_a__U32,id_a__isNatIListKind => false | id_a__U32,id_a__U12 => false | id_a__U32,id_U31 => false | id_a__U32,id_a__length => false | id_a__U32,id_a__U42 => false | id_a__U32,id_U53 => false | id_a__U32,id_and => false | id_a__U32,id_cons => false | id_a__U32,id_U11 => false | id_a__U32,id_a__U51 => false | id_a__U32,id_a__U22 => false | id_a__U32,id_U42 => false | id_a__U32,id_length => false | id_a__U32,id_tt => false | id_a__U32,id_U22 => false | id_a__U32,id_s => false | id_a__U32,id_a__U41 => false | id_a__U32,id_U52 => false | id_a__U32,id_nil => false | id_a__U32,id_zeros => false | id_a__U32,id_isNatList => false | id_a__U32,id_a__U53 => false | id_a__U32,id_a__U31 => false | id_a__U32,id_isNatIList => false | id_a__U32,id_a__isNatKind => false | id_a__U32,id_a__isNatList => false | id_a__U32,id_U32 => false | id_a__U32,id_mark => false | id_a__U32,id_a__U43 => false | id_a__U32,id_U61 => false | id_a__U32,id_isNat => false | id_U51,id_a__zeros => true | id_U51,id_isNatKind => true | id_U51,id_a__isNatIList => true | id_U51,id_a__U21 => true | id_U51,id_U41 => true | id_U51,id_a__and => true | id_U51,id_a__U11 => true | id_U51,id_U21 => true | id_U51,id_a__U61 => true | id_U51,id_a__U32 => true | id_U51,id_U51 => true | id_U51,id_isNatIListKind => false | id_U51,id_0 => false | id_U51,id_U12 => false | id_U51,id_a__U52 => false | id_U51,id_a__isNat => false | id_U51,id_U43 => false | id_U51,id_a__isNatIListKind => false | id_U51,id_a__U12 => false | id_U51,id_U31 => false | id_U51,id_a__length => false | id_U51,id_a__U42 => false | id_U51,id_U53 => false | id_U51,id_and => false | id_U51,id_cons => false | id_U51,id_U11 => false | id_U51,id_a__U51 => false | id_U51,id_a__U22 => false | id_U51,id_U42 => false | id_U51,id_length => false | id_U51,id_tt => false | id_U51,id_U22 => false | id_U51,id_s => false | id_U51,id_a__U41 => false | id_U51,id_U52 => false | id_U51,id_nil => false | id_U51,id_zeros => false | id_U51,id_isNatList => false | id_U51,id_a__U53 => false | id_U51,id_a__U31 => false | id_U51,id_isNatIList => false | id_U51,id_a__isNatKind => false | id_U51,id_a__isNatList => false | id_U51,id_U32 => false | id_U51,id_mark => false | id_U51,id_a__U43 => false | id_U51,id_U61 => false | id_U51,id_isNat => false | id_isNatIListKind,id_a__zeros => true | id_isNatIListKind,id_isNatKind => true | id_isNatIListKind,id_a__isNatIList => true | id_isNatIListKind,id_a__U21 => true | id_isNatIListKind,id_U41 => true | id_isNatIListKind,id_a__and => true | id_isNatIListKind,id_a__U11 => true | id_isNatIListKind,id_U21 => true | id_isNatIListKind,id_a__U61 => true | id_isNatIListKind,id_a__U32 => true | id_isNatIListKind,id_U51 => true | id_isNatIListKind,id_isNatIListKind => true | id_isNatIListKind,id_0 => false | id_isNatIListKind,id_U12 => false | id_isNatIListKind,id_a__U52 => false | id_isNatIListKind,id_a__isNat => false | id_isNatIListKind,id_U43 => false | id_isNatIListKind,id_a__isNatIListKind => false | id_isNatIListKind,id_a__U12 => false | id_isNatIListKind,id_U31 => false | id_isNatIListKind,id_a__length => false | id_isNatIListKind,id_a__U42 => false | id_isNatIListKind,id_U53 => false | id_isNatIListKind,id_and => false | id_isNatIListKind,id_cons => false | id_isNatIListKind,id_U11 => false | id_isNatIListKind,id_a__U51 => false | id_isNatIListKind,id_a__U22 => false | id_isNatIListKind,id_U42 => false | id_isNatIListKind,id_length => false | id_isNatIListKind,id_tt => false | id_isNatIListKind,id_U22 => false | id_isNatIListKind,id_s => false | id_isNatIListKind,id_a__U41 => false | id_isNatIListKind,id_U52 => false | id_isNatIListKind,id_nil => false | id_isNatIListKind,id_zeros => false | id_isNatIListKind,id_isNatList => false | id_isNatIListKind,id_a__U53 => false | id_isNatIListKind,id_a__U31 => false | id_isNatIListKind,id_isNatIList => false | id_isNatIListKind,id_a__isNatKind => false | id_isNatIListKind,id_a__isNatList => false | id_isNatIListKind,id_U32 => false | id_isNatIListKind,id_mark => false | id_isNatIListKind,id_a__U43 => false | id_isNatIListKind,id_U61 => false | id_isNatIListKind,id_isNat => false | id_0,id_a__zeros => true | id_0,id_isNatKind => true | id_0,id_a__isNatIList => true | id_0,id_a__U21 => true | id_0,id_U41 => true | id_0,id_a__and => true | id_0,id_a__U11 => true | id_0,id_U21 => true | id_0,id_a__U61 => true | id_0,id_a__U32 => true | id_0,id_U51 => true | id_0,id_isNatIListKind => true | id_0,id_0 => true | id_0,id_U12 => false | id_0,id_a__U52 => false | id_0,id_a__isNat => false | id_0,id_U43 => false | id_0,id_a__isNatIListKind => false | id_0,id_a__U12 => false | id_0,id_U31 => false | id_0,id_a__length => false | id_0,id_a__U42 => false | id_0,id_U53 => false | id_0,id_and => false | id_0,id_cons => false | id_0,id_U11 => false | id_0,id_a__U51 => false | id_0,id_a__U22 => false | id_0,id_U42 => false | id_0,id_length => false | id_0,id_tt => false | id_0,id_U22 => false | id_0,id_s => false | id_0,id_a__U41 => false | id_0,id_U52 => false | id_0,id_nil => false | id_0,id_zeros => false | id_0,id_isNatList => false | id_0,id_a__U53 => false | id_0,id_a__U31 => false | id_0,id_isNatIList => false | id_0,id_a__isNatKind => false | id_0,id_a__isNatList => false | id_0,id_U32 => false | id_0,id_mark => false | id_0,id_a__U43 => false | id_0,id_U61 => false | id_0,id_isNat => false | id_U12,id_a__zeros => true | id_U12,id_isNatKind => true | id_U12,id_a__isNatIList => true | id_U12,id_a__U21 => true | id_U12,id_U41 => true | id_U12,id_a__and => true | id_U12,id_a__U11 => true | id_U12,id_U21 => true | id_U12,id_a__U61 => true | id_U12,id_a__U32 => true | id_U12,id_U51 => true | id_U12,id_isNatIListKind => true | id_U12,id_0 => true | id_U12,id_U12 => true | id_U12,id_a__U52 => false | id_U12,id_a__isNat => false | id_U12,id_U43 => false | id_U12,id_a__isNatIListKind => false | id_U12,id_a__U12 => false | id_U12,id_U31 => false | id_U12,id_a__length => false | id_U12,id_a__U42 => false | id_U12,id_U53 => false | id_U12,id_and => false | id_U12,id_cons => false | id_U12,id_U11 => false | id_U12,id_a__U51 => false | id_U12,id_a__U22 => false | id_U12,id_U42 => false | id_U12,id_length => false | id_U12,id_tt => false | id_U12,id_U22 => false | id_U12,id_s => false | id_U12,id_a__U41 => false | id_U12,id_U52 => false | id_U12,id_nil => false | id_U12,id_zeros => false | id_U12,id_isNatList => false | id_U12,id_a__U53 => false | id_U12,id_a__U31 => false | id_U12,id_isNatIList => false | id_U12,id_a__isNatKind => false | id_U12,id_a__isNatList => false | id_U12,id_U32 => false | id_U12,id_mark => false | id_U12,id_a__U43 => false | id_U12,id_U61 => false | id_U12,id_isNat => false | id_a__U52,id_a__zeros => true | id_a__U52,id_isNatKind => true | id_a__U52,id_a__isNatIList => true | id_a__U52,id_a__U21 => true | id_a__U52,id_U41 => true | id_a__U52,id_a__and => true | id_a__U52,id_a__U11 => true | id_a__U52,id_U21 => true | id_a__U52,id_a__U61 => true | id_a__U52,id_a__U32 => true | id_a__U52,id_U51 => true | id_a__U52,id_isNatIListKind => true | id_a__U52,id_0 => true | id_a__U52,id_U12 => true | id_a__U52,id_a__U52 => true | id_a__U52,id_a__isNat => false | id_a__U52,id_U43 => false | id_a__U52,id_a__isNatIListKind => false | id_a__U52,id_a__U12 => false | id_a__U52,id_U31 => false | id_a__U52,id_a__length => false | id_a__U52,id_a__U42 => false | id_a__U52,id_U53 => false | id_a__U52,id_and => false | id_a__U52,id_cons => false | id_a__U52,id_U11 => false | id_a__U52,id_a__U51 => false | id_a__U52,id_a__U22 => false | id_a__U52,id_U42 => false | id_a__U52,id_length => false | id_a__U52,id_tt => false | id_a__U52,id_U22 => false | id_a__U52,id_s => false | id_a__U52,id_a__U41 => false | id_a__U52,id_U52 => false | id_a__U52,id_nil => false | id_a__U52,id_zeros => false | id_a__U52,id_isNatList => false | id_a__U52,id_a__U53 => false | id_a__U52,id_a__U31 => false | id_a__U52,id_isNatIList => false | id_a__U52,id_a__isNatKind => false | id_a__U52,id_a__isNatList => false | id_a__U52,id_U32 => false | id_a__U52,id_mark => false | id_a__U52,id_a__U43 => false | id_a__U52,id_U61 => false | id_a__U52,id_isNat => false | id_a__isNat,id_a__zeros => true | id_a__isNat,id_isNatKind => true | id_a__isNat,id_a__isNatIList => true | id_a__isNat,id_a__U21 => true | id_a__isNat,id_U41 => true | id_a__isNat,id_a__and => true | id_a__isNat,id_a__U11 => true | id_a__isNat,id_U21 => true | id_a__isNat,id_a__U61 => true | id_a__isNat,id_a__U32 => true | id_a__isNat,id_U51 => true | id_a__isNat,id_isNatIListKind => true | id_a__isNat,id_0 => true | id_a__isNat,id_U12 => true | id_a__isNat,id_a__U52 => true | id_a__isNat,id_a__isNat => true | id_a__isNat,id_U43 => false | id_a__isNat,id_a__isNatIListKind => false | id_a__isNat,id_a__U12 => false | id_a__isNat,id_U31 => false | id_a__isNat,id_a__length => false | id_a__isNat,id_a__U42 => false | id_a__isNat,id_U53 => false | id_a__isNat,id_and => false | id_a__isNat,id_cons => false | id_a__isNat,id_U11 => false | id_a__isNat,id_a__U51 => false | id_a__isNat,id_a__U22 => false | id_a__isNat,id_U42 => false | id_a__isNat,id_length => false | id_a__isNat,id_tt => false | id_a__isNat,id_U22 => false | id_a__isNat,id_s => false | id_a__isNat,id_a__U41 => false | id_a__isNat,id_U52 => false | id_a__isNat,id_nil => false | id_a__isNat,id_zeros => false | id_a__isNat,id_isNatList => false | id_a__isNat,id_a__U53 => false | id_a__isNat,id_a__U31 => false | id_a__isNat,id_isNatIList => false | id_a__isNat,id_a__isNatKind => false | id_a__isNat,id_a__isNatList => false | id_a__isNat,id_U32 => false | id_a__isNat,id_mark => false | id_a__isNat,id_a__U43 => false | id_a__isNat,id_U61 => false | id_a__isNat,id_isNat => false | id_U43,id_a__zeros => true | id_U43,id_isNatKind => true | id_U43,id_a__isNatIList => true | id_U43,id_a__U21 => true | id_U43,id_U41 => true | id_U43,id_a__and => true | id_U43,id_a__U11 => true | id_U43,id_U21 => true | id_U43,id_a__U61 => true | id_U43,id_a__U32 => true | id_U43,id_U51 => true | id_U43,id_isNatIListKind => true | id_U43,id_0 => true | id_U43,id_U12 => true | id_U43,id_a__U52 => true | id_U43,id_a__isNat => true | id_U43,id_U43 => true | id_U43,id_a__isNatIListKind => false | id_U43,id_a__U12 => false | id_U43,id_U31 => false | id_U43,id_a__length => false | id_U43,id_a__U42 => false | id_U43,id_U53 => false | id_U43,id_and => false | id_U43,id_cons => false | id_U43,id_U11 => false | id_U43,id_a__U51 => false | id_U43,id_a__U22 => false | id_U43,id_U42 => false | id_U43,id_length => false | id_U43,id_tt => false | id_U43,id_U22 => false | id_U43,id_s => false | id_U43,id_a__U41 => false | id_U43,id_U52 => false | id_U43,id_nil => false | id_U43,id_zeros => false | id_U43,id_isNatList => false | id_U43,id_a__U53 => false | id_U43,id_a__U31 => false | id_U43,id_isNatIList => false | id_U43,id_a__isNatKind => false | id_U43,id_a__isNatList => false | id_U43,id_U32 => false | id_U43,id_mark => false | id_U43,id_a__U43 => false | id_U43,id_U61 => false | id_U43,id_isNat => false | id_a__isNatIListKind,id_a__zeros => true | id_a__isNatIListKind,id_isNatKind => true | id_a__isNatIListKind,id_a__isNatIList => true | id_a__isNatIListKind,id_a__U21 => true | id_a__isNatIListKind,id_U41 => true | id_a__isNatIListKind,id_a__and => true | id_a__isNatIListKind,id_a__U11 => true | id_a__isNatIListKind,id_U21 => true | id_a__isNatIListKind,id_a__U61 => true | id_a__isNatIListKind,id_a__U32 => true | id_a__isNatIListKind,id_U51 => true | id_a__isNatIListKind,id_isNatIListKind => true | id_a__isNatIListKind,id_0 => true | id_a__isNatIListKind,id_U12 => true | id_a__isNatIListKind,id_a__U52 => true | id_a__isNatIListKind,id_a__isNat => true | id_a__isNatIListKind,id_U43 => true | id_a__isNatIListKind,id_a__isNatIListKind => true | id_a__isNatIListKind,id_a__U12 => false | id_a__isNatIListKind,id_U31 => false | id_a__isNatIListKind,id_a__length => false | id_a__isNatIListKind,id_a__U42 => false | id_a__isNatIListKind,id_U53 => false | id_a__isNatIListKind,id_and => false | id_a__isNatIListKind,id_cons => false | id_a__isNatIListKind,id_U11 => false | id_a__isNatIListKind,id_a__U51 => false | id_a__isNatIListKind,id_a__U22 => false | id_a__isNatIListKind,id_U42 => false | id_a__isNatIListKind,id_length => false | id_a__isNatIListKind,id_tt => false | id_a__isNatIListKind,id_U22 => false | id_a__isNatIListKind,id_s => false | id_a__isNatIListKind,id_a__U41 => false | id_a__isNatIListKind,id_U52 => false | id_a__isNatIListKind,id_nil => false | id_a__isNatIListKind,id_zeros => false | id_a__isNatIListKind,id_isNatList => false | id_a__isNatIListKind,id_a__U53 => false | id_a__isNatIListKind,id_a__U31 => false | id_a__isNatIListKind,id_isNatIList => false | id_a__isNatIListKind,id_a__isNatKind => false | id_a__isNatIListKind,id_a__isNatList => false | id_a__isNatIListKind,id_U32 => false | id_a__isNatIListKind,id_mark => false | id_a__isNatIListKind,id_a__U43 => false | id_a__isNatIListKind,id_U61 => false | id_a__isNatIListKind,id_isNat => false | id_a__U12,id_a__zeros => true | id_a__U12,id_isNatKind => true | id_a__U12,id_a__isNatIList => true | id_a__U12,id_a__U21 => true | id_a__U12,id_U41 => true | id_a__U12,id_a__and => true | id_a__U12,id_a__U11 => true | id_a__U12,id_U21 => true | id_a__U12,id_a__U61 => true | id_a__U12,id_a__U32 => true | id_a__U12,id_U51 => true | id_a__U12,id_isNatIListKind => true | id_a__U12,id_0 => true | id_a__U12,id_U12 => true | id_a__U12,id_a__U52 => true | id_a__U12,id_a__isNat => true | id_a__U12,id_U43 => true | id_a__U12,id_a__isNatIListKind => true | id_a__U12,id_a__U12 => true | id_a__U12,id_U31 => false | id_a__U12,id_a__length => false | id_a__U12,id_a__U42 => false | id_a__U12,id_U53 => false | id_a__U12,id_and => false | id_a__U12,id_cons => false | id_a__U12,id_U11 => false | id_a__U12,id_a__U51 => false | id_a__U12,id_a__U22 => false | id_a__U12,id_U42 => false | id_a__U12,id_length => false | id_a__U12,id_tt => false | id_a__U12,id_U22 => false | id_a__U12,id_s => false | id_a__U12,id_a__U41 => false | id_a__U12,id_U52 => false | id_a__U12,id_nil => false | id_a__U12,id_zeros => false | id_a__U12,id_isNatList => false | id_a__U12,id_a__U53 => false | id_a__U12,id_a__U31 => false | id_a__U12,id_isNatIList => false | id_a__U12,id_a__isNatKind => false | id_a__U12,id_a__isNatList => false | id_a__U12,id_U32 => false | id_a__U12,id_mark => false | id_a__U12,id_a__U43 => false | id_a__U12,id_U61 => false | id_a__U12,id_isNat => false | id_U31,id_a__zeros => true | id_U31,id_isNatKind => true | id_U31,id_a__isNatIList => true | id_U31,id_a__U21 => true | id_U31,id_U41 => true | id_U31,id_a__and => true | id_U31,id_a__U11 => true | id_U31,id_U21 => true | id_U31,id_a__U61 => true | id_U31,id_a__U32 => true | id_U31,id_U51 => true | id_U31,id_isNatIListKind => true | id_U31,id_0 => true | id_U31,id_U12 => true | id_U31,id_a__U52 => true | id_U31,id_a__isNat => true | id_U31,id_U43 => true | id_U31,id_a__isNatIListKind => true | id_U31,id_a__U12 => true | id_U31,id_U31 => true | id_U31,id_a__length => false | id_U31,id_a__U42 => false | id_U31,id_U53 => false | id_U31,id_and => false | id_U31,id_cons => false | id_U31,id_U11 => false | id_U31,id_a__U51 => false | id_U31,id_a__U22 => false | id_U31,id_U42 => false | id_U31,id_length => false | id_U31,id_tt => false | id_U31,id_U22 => false | id_U31,id_s => false | id_U31,id_a__U41 => false | id_U31,id_U52 => false | id_U31,id_nil => false | id_U31,id_zeros => false | id_U31,id_isNatList => false | id_U31,id_a__U53 => false | id_U31,id_a__U31 => false | id_U31,id_isNatIList => false | id_U31,id_a__isNatKind => false | id_U31,id_a__isNatList => false | id_U31,id_U32 => false | id_U31,id_mark => false | id_U31,id_a__U43 => false | id_U31,id_U61 => false | id_U31,id_isNat => false | id_a__length,id_a__zeros => true | id_a__length,id_isNatKind => true | id_a__length,id_a__isNatIList => true | id_a__length,id_a__U21 => true | id_a__length,id_U41 => true | id_a__length,id_a__and => true | id_a__length,id_a__U11 => true | id_a__length,id_U21 => true | id_a__length,id_a__U61 => true | id_a__length,id_a__U32 => true | id_a__length,id_U51 => true | id_a__length,id_isNatIListKind => true | id_a__length,id_0 => true | id_a__length,id_U12 => true | id_a__length,id_a__U52 => true | id_a__length,id_a__isNat => true | id_a__length,id_U43 => true | id_a__length,id_a__isNatIListKind => true | id_a__length,id_a__U12 => true | id_a__length,id_U31 => true | id_a__length,id_a__length => true | id_a__length,id_a__U42 => false | id_a__length,id_U53 => false | id_a__length,id_and => false | id_a__length,id_cons => false | id_a__length,id_U11 => false | id_a__length,id_a__U51 => false | id_a__length,id_a__U22 => false | id_a__length,id_U42 => false | id_a__length,id_length => false | id_a__length,id_tt => false | id_a__length,id_U22 => false | id_a__length,id_s => false | id_a__length,id_a__U41 => false | id_a__length,id_U52 => false | id_a__length,id_nil => false | id_a__length,id_zeros => false | id_a__length,id_isNatList => false | id_a__length,id_a__U53 => false | id_a__length,id_a__U31 => false | id_a__length,id_isNatIList => false | id_a__length,id_a__isNatKind => false | id_a__length,id_a__isNatList => false | id_a__length,id_U32 => false | id_a__length,id_mark => false | id_a__length,id_a__U43 => false | id_a__length,id_U61 => false | id_a__length,id_isNat => false | id_a__U42,id_a__zeros => true | id_a__U42,id_isNatKind => true | id_a__U42,id_a__isNatIList => true | id_a__U42,id_a__U21 => true | id_a__U42,id_U41 => true | id_a__U42,id_a__and => true | id_a__U42,id_a__U11 => true | id_a__U42,id_U21 => true | id_a__U42,id_a__U61 => true | id_a__U42,id_a__U32 => true | id_a__U42,id_U51 => true | id_a__U42,id_isNatIListKind => true | id_a__U42,id_0 => true | id_a__U42,id_U12 => true | id_a__U42,id_a__U52 => true | id_a__U42,id_a__isNat => true | id_a__U42,id_U43 => true | id_a__U42,id_a__isNatIListKind => true | id_a__U42,id_a__U12 => true | id_a__U42,id_U31 => true | id_a__U42,id_a__length => true | id_a__U42,id_a__U42 => true | id_a__U42,id_U53 => false | id_a__U42,id_and => false | id_a__U42,id_cons => false | id_a__U42,id_U11 => false | id_a__U42,id_a__U51 => false | id_a__U42,id_a__U22 => false | id_a__U42,id_U42 => false | id_a__U42,id_length => false | id_a__U42,id_tt => false | id_a__U42,id_U22 => false | id_a__U42,id_s => false | id_a__U42,id_a__U41 => false | id_a__U42,id_U52 => false | id_a__U42,id_nil => false | id_a__U42,id_zeros => false | id_a__U42,id_isNatList => false | id_a__U42,id_a__U53 => false | id_a__U42,id_a__U31 => false | id_a__U42,id_isNatIList => false | id_a__U42,id_a__isNatKind => false | id_a__U42,id_a__isNatList => false | id_a__U42,id_U32 => false | id_a__U42,id_mark => false | id_a__U42,id_a__U43 => false | id_a__U42,id_U61 => false | id_a__U42,id_isNat => false | id_U53,id_a__zeros => true | id_U53,id_isNatKind => true | id_U53,id_a__isNatIList => true | id_U53,id_a__U21 => true | id_U53,id_U41 => true | id_U53,id_a__and => true | id_U53,id_a__U11 => true | id_U53,id_U21 => true | id_U53,id_a__U61 => true | id_U53,id_a__U32 => true | id_U53,id_U51 => true | id_U53,id_isNatIListKind => true | id_U53,id_0 => true | id_U53,id_U12 => true | id_U53,id_a__U52 => true | id_U53,id_a__isNat => true | id_U53,id_U43 => true | id_U53,id_a__isNatIListKind => true | id_U53,id_a__U12 => true | id_U53,id_U31 => true | id_U53,id_a__length => true | id_U53,id_a__U42 => true | id_U53,id_U53 => true | id_U53,id_and => false | id_U53,id_cons => false | id_U53,id_U11 => false | id_U53,id_a__U51 => false | id_U53,id_a__U22 => false | id_U53,id_U42 => false | id_U53,id_length => false | id_U53,id_tt => false | id_U53,id_U22 => false | id_U53,id_s => false | id_U53,id_a__U41 => false | id_U53,id_U52 => false | id_U53,id_nil => false | id_U53,id_zeros => false | id_U53,id_isNatList => false | id_U53,id_a__U53 => false | id_U53,id_a__U31 => false | id_U53,id_isNatIList => false | id_U53,id_a__isNatKind => false | id_U53,id_a__isNatList => false | id_U53,id_U32 => false | id_U53,id_mark => false | id_U53,id_a__U43 => false | id_U53,id_U61 => false | id_U53,id_isNat => false | id_and,id_a__zeros => true | id_and,id_isNatKind => true | id_and,id_a__isNatIList => true | id_and,id_a__U21 => true | id_and,id_U41 => true | id_and,id_a__and => true | id_and,id_a__U11 => true | id_and,id_U21 => true | id_and,id_a__U61 => true | id_and,id_a__U32 => true | id_and,id_U51 => true | id_and,id_isNatIListKind => true | id_and,id_0 => true | id_and,id_U12 => true | id_and,id_a__U52 => true | id_and,id_a__isNat => true | id_and,id_U43 => true | id_and,id_a__isNatIListKind => true | id_and,id_a__U12 => true | id_and,id_U31 => true | id_and,id_a__length => true | id_and,id_a__U42 => true | id_and,id_U53 => true | id_and,id_and => true | id_and,id_cons => false | id_and,id_U11 => false | id_and,id_a__U51 => false | id_and,id_a__U22 => false | id_and,id_U42 => false | id_and,id_length => false | id_and,id_tt => false | id_and,id_U22 => false | id_and,id_s => false | id_and,id_a__U41 => false | id_and,id_U52 => false | id_and,id_nil => false | id_and,id_zeros => false | id_and,id_isNatList => false | id_and,id_a__U53 => false | id_and,id_a__U31 => false | id_and,id_isNatIList => false | id_and,id_a__isNatKind => false | id_and,id_a__isNatList => false | id_and,id_U32 => false | id_and,id_mark => false | id_and,id_a__U43 => false | id_and,id_U61 => false | id_and,id_isNat => false | id_cons,id_a__zeros => true | id_cons,id_isNatKind => true | id_cons,id_a__isNatIList => true | id_cons,id_a__U21 => true | id_cons,id_U41 => true | id_cons,id_a__and => true | id_cons,id_a__U11 => true | id_cons,id_U21 => true | id_cons,id_a__U61 => true | id_cons,id_a__U32 => true | id_cons,id_U51 => true | id_cons,id_isNatIListKind => true | id_cons,id_0 => true | id_cons,id_U12 => true | id_cons,id_a__U52 => true | id_cons,id_a__isNat => true | id_cons,id_U43 => true | id_cons,id_a__isNatIListKind => true | id_cons,id_a__U12 => true | id_cons,id_U31 => true | id_cons,id_a__length => true | id_cons,id_a__U42 => true | id_cons,id_U53 => true | id_cons,id_and => true | id_cons,id_cons => true | id_cons,id_U11 => false | id_cons,id_a__U51 => false | id_cons,id_a__U22 => false | id_cons,id_U42 => false | id_cons,id_length => false | id_cons,id_tt => false | id_cons,id_U22 => false | id_cons,id_s => false | id_cons,id_a__U41 => false | id_cons,id_U52 => false | id_cons,id_nil => false | id_cons,id_zeros => false | id_cons,id_isNatList => false | id_cons,id_a__U53 => false | id_cons,id_a__U31 => false | id_cons,id_isNatIList => false | id_cons,id_a__isNatKind => false | id_cons,id_a__isNatList => false | id_cons,id_U32 => false | id_cons,id_mark => false | id_cons,id_a__U43 => false | id_cons,id_U61 => false | id_cons,id_isNat => false | id_U11,id_a__zeros => true | id_U11,id_isNatKind => true | id_U11,id_a__isNatIList => true | id_U11,id_a__U21 => true | id_U11,id_U41 => true | id_U11,id_a__and => true | id_U11,id_a__U11 => true | id_U11,id_U21 => true | id_U11,id_a__U61 => true | id_U11,id_a__U32 => true | id_U11,id_U51 => true | id_U11,id_isNatIListKind => true | id_U11,id_0 => true | id_U11,id_U12 => true | id_U11,id_a__U52 => true | id_U11,id_a__isNat => true | id_U11,id_U43 => true | id_U11,id_a__isNatIListKind => true | id_U11,id_a__U12 => true | id_U11,id_U31 => true | id_U11,id_a__length => true | id_U11,id_a__U42 => true | id_U11,id_U53 => true | id_U11,id_and => true | id_U11,id_cons => true | id_U11,id_U11 => true | id_U11,id_a__U51 => false | id_U11,id_a__U22 => false | id_U11,id_U42 => false | id_U11,id_length => false | id_U11,id_tt => false | id_U11,id_U22 => false | id_U11,id_s => false | id_U11,id_a__U41 => false | id_U11,id_U52 => false | id_U11,id_nil => false | id_U11,id_zeros => false | id_U11,id_isNatList => false | id_U11,id_a__U53 => false | id_U11,id_a__U31 => false | id_U11,id_isNatIList => false | id_U11,id_a__isNatKind => false | id_U11,id_a__isNatList => false | id_U11,id_U32 => false | id_U11,id_mark => false | id_U11,id_a__U43 => false | id_U11,id_U61 => false | id_U11,id_isNat => false | id_a__U51,id_a__zeros => true | id_a__U51,id_isNatKind => true | id_a__U51,id_a__isNatIList => true | id_a__U51,id_a__U21 => true | id_a__U51,id_U41 => true | id_a__U51,id_a__and => true | id_a__U51,id_a__U11 => true | id_a__U51,id_U21 => true | id_a__U51,id_a__U61 => true | id_a__U51,id_a__U32 => true | id_a__U51,id_U51 => true | id_a__U51,id_isNatIListKind => true | id_a__U51,id_0 => true | id_a__U51,id_U12 => true | id_a__U51,id_a__U52 => true | id_a__U51,id_a__isNat => true | id_a__U51,id_U43 => true | id_a__U51,id_a__isNatIListKind => true | id_a__U51,id_a__U12 => true | id_a__U51,id_U31 => true | id_a__U51,id_a__length => true | id_a__U51,id_a__U42 => true | id_a__U51,id_U53 => true | id_a__U51,id_and => true | id_a__U51,id_cons => true | id_a__U51,id_U11 => true | id_a__U51,id_a__U51 => true | id_a__U51,id_a__U22 => false | id_a__U51,id_U42 => false | id_a__U51,id_length => false | id_a__U51,id_tt => false | id_a__U51,id_U22 => false | id_a__U51,id_s => false | id_a__U51,id_a__U41 => false | id_a__U51,id_U52 => false | id_a__U51,id_nil => false | id_a__U51,id_zeros => false | id_a__U51,id_isNatList => false | id_a__U51,id_a__U53 => false | id_a__U51,id_a__U31 => false | id_a__U51,id_isNatIList => false | id_a__U51,id_a__isNatKind => false | id_a__U51,id_a__isNatList => false | id_a__U51,id_U32 => false | id_a__U51,id_mark => false | id_a__U51,id_a__U43 => false | id_a__U51,id_U61 => false | id_a__U51,id_isNat => false | id_a__U22,id_a__zeros => true | id_a__U22,id_isNatKind => true | id_a__U22,id_a__isNatIList => true | id_a__U22,id_a__U21 => true | id_a__U22,id_U41 => true | id_a__U22,id_a__and => true | id_a__U22,id_a__U11 => true | id_a__U22,id_U21 => true | id_a__U22,id_a__U61 => true | id_a__U22,id_a__U32 => true | id_a__U22,id_U51 => true | id_a__U22,id_isNatIListKind => true | id_a__U22,id_0 => true | id_a__U22,id_U12 => true | id_a__U22,id_a__U52 => true | id_a__U22,id_a__isNat => true | id_a__U22,id_U43 => true | id_a__U22,id_a__isNatIListKind => true | id_a__U22,id_a__U12 => true | id_a__U22,id_U31 => true | id_a__U22,id_a__length => true | id_a__U22,id_a__U42 => true | id_a__U22,id_U53 => true | id_a__U22,id_and => true | id_a__U22,id_cons => true | id_a__U22,id_U11 => true | id_a__U22,id_a__U51 => true | id_a__U22,id_a__U22 => true | id_a__U22,id_U42 => false | id_a__U22,id_length => false | id_a__U22,id_tt => false | id_a__U22,id_U22 => false | id_a__U22,id_s => false | id_a__U22,id_a__U41 => false | id_a__U22,id_U52 => false | id_a__U22,id_nil => false | id_a__U22,id_zeros => false | id_a__U22,id_isNatList => false | id_a__U22,id_a__U53 => false | id_a__U22,id_a__U31 => false | id_a__U22,id_isNatIList => false | id_a__U22,id_a__isNatKind => false | id_a__U22,id_a__isNatList => false | id_a__U22,id_U32 => false | id_a__U22,id_mark => false | id_a__U22,id_a__U43 => false | id_a__U22,id_U61 => false | id_a__U22,id_isNat => false | id_U42,id_a__zeros => true | id_U42,id_isNatKind => true | id_U42,id_a__isNatIList => true | id_U42,id_a__U21 => true | id_U42,id_U41 => true | id_U42,id_a__and => true | id_U42,id_a__U11 => true | id_U42,id_U21 => true | id_U42,id_a__U61 => true | id_U42,id_a__U32 => true | id_U42,id_U51 => true | id_U42,id_isNatIListKind => true | id_U42,id_0 => true | id_U42,id_U12 => true | id_U42,id_a__U52 => true | id_U42,id_a__isNat => true | id_U42,id_U43 => true | id_U42,id_a__isNatIListKind => true | id_U42,id_a__U12 => true | id_U42,id_U31 => true | id_U42,id_a__length => true | id_U42,id_a__U42 => true | id_U42,id_U53 => true | id_U42,id_and => true | id_U42,id_cons => true | id_U42,id_U11 => true | id_U42,id_a__U51 => true | id_U42,id_a__U22 => true | id_U42,id_U42 => true | id_U42,id_length => false | id_U42,id_tt => false | id_U42,id_U22 => false | id_U42,id_s => false | id_U42,id_a__U41 => false | id_U42,id_U52 => false | id_U42,id_nil => false | id_U42,id_zeros => false | id_U42,id_isNatList => false | id_U42,id_a__U53 => false | id_U42,id_a__U31 => false | id_U42,id_isNatIList => false | id_U42,id_a__isNatKind => false | id_U42,id_a__isNatList => false | id_U42,id_U32 => false | id_U42,id_mark => false | id_U42,id_a__U43 => false | id_U42,id_U61 => false | id_U42,id_isNat => false | id_length,id_a__zeros => true | id_length,id_isNatKind => true | id_length,id_a__isNatIList => true | id_length,id_a__U21 => true | id_length,id_U41 => true | id_length,id_a__and => true | id_length,id_a__U11 => true | id_length,id_U21 => true | id_length,id_a__U61 => true | id_length,id_a__U32 => true | id_length,id_U51 => true | id_length,id_isNatIListKind => true | id_length,id_0 => true | id_length,id_U12 => true | id_length,id_a__U52 => true | id_length,id_a__isNat => true | id_length,id_U43 => true | id_length,id_a__isNatIListKind => true | id_length,id_a__U12 => true | id_length,id_U31 => true | id_length,id_a__length => true | id_length,id_a__U42 => true | id_length,id_U53 => true | id_length,id_and => true | id_length,id_cons => true | id_length,id_U11 => true | id_length,id_a__U51 => true | id_length,id_a__U22 => true | id_length,id_U42 => true | id_length,id_length => true | id_length,id_tt => false | id_length,id_U22 => false | id_length,id_s => false | id_length,id_a__U41 => false | id_length,id_U52 => false | id_length,id_nil => false | id_length,id_zeros => false | id_length,id_isNatList => false | id_length,id_a__U53 => false | id_length,id_a__U31 => false | id_length,id_isNatIList => false | id_length,id_a__isNatKind => false | id_length,id_a__isNatList => false | id_length,id_U32 => false | id_length,id_mark => false | id_length,id_a__U43 => false | id_length,id_U61 => false | id_length,id_isNat => false | id_tt,id_a__zeros => true | id_tt,id_isNatKind => true | id_tt,id_a__isNatIList => true | id_tt,id_a__U21 => true | id_tt,id_U41 => true | id_tt,id_a__and => true | id_tt,id_a__U11 => true | id_tt,id_U21 => true | id_tt,id_a__U61 => true | id_tt,id_a__U32 => true | id_tt,id_U51 => true | id_tt,id_isNatIListKind => true | id_tt,id_0 => true | id_tt,id_U12 => true | id_tt,id_a__U52 => true | id_tt,id_a__isNat => true | id_tt,id_U43 => true | id_tt,id_a__isNatIListKind => true | id_tt,id_a__U12 => true | id_tt,id_U31 => true | id_tt,id_a__length => true | id_tt,id_a__U42 => true | id_tt,id_U53 => true | id_tt,id_and => true | id_tt,id_cons => true | id_tt,id_U11 => true | id_tt,id_a__U51 => true | id_tt,id_a__U22 => true | id_tt,id_U42 => true | id_tt,id_length => true | id_tt,id_tt => true | id_tt,id_U22 => false | id_tt,id_s => false | id_tt,id_a__U41 => false | id_tt,id_U52 => false | id_tt,id_nil => false | id_tt,id_zeros => false | id_tt,id_isNatList => false | id_tt,id_a__U53 => false | id_tt,id_a__U31 => false | id_tt,id_isNatIList => false | id_tt,id_a__isNatKind => false | id_tt,id_a__isNatList => false | id_tt,id_U32 => false | id_tt,id_mark => false | id_tt,id_a__U43 => false | id_tt,id_U61 => false | id_tt,id_isNat => false | id_U22,id_a__zeros => true | id_U22,id_isNatKind => true | id_U22,id_a__isNatIList => true | id_U22,id_a__U21 => true | id_U22,id_U41 => true | id_U22,id_a__and => true | id_U22,id_a__U11 => true | id_U22,id_U21 => true | id_U22,id_a__U61 => true | id_U22,id_a__U32 => true | id_U22,id_U51 => true | id_U22,id_isNatIListKind => true | id_U22,id_0 => true | id_U22,id_U12 => true | id_U22,id_a__U52 => true | id_U22,id_a__isNat => true | id_U22,id_U43 => true | id_U22,id_a__isNatIListKind => true | id_U22,id_a__U12 => true | id_U22,id_U31 => true | id_U22,id_a__length => true | id_U22,id_a__U42 => true | id_U22,id_U53 => true | id_U22,id_and => true | id_U22,id_cons => true | id_U22,id_U11 => true | id_U22,id_a__U51 => true | id_U22,id_a__U22 => true | id_U22,id_U42 => true | id_U22,id_length => true | id_U22,id_tt => true | id_U22,id_U22 => true | id_U22,id_s => false | id_U22,id_a__U41 => false | id_U22,id_U52 => false | id_U22,id_nil => false | id_U22,id_zeros => false | id_U22,id_isNatList => false | id_U22,id_a__U53 => false | id_U22,id_a__U31 => false | id_U22,id_isNatIList => false | id_U22,id_a__isNatKind => false | id_U22,id_a__isNatList => false | id_U22,id_U32 => false | id_U22,id_mark => false | id_U22,id_a__U43 => false | id_U22,id_U61 => false | id_U22,id_isNat => false | id_s,id_a__zeros => true | id_s,id_isNatKind => true | id_s,id_a__isNatIList => true | id_s,id_a__U21 => true | id_s,id_U41 => true | id_s,id_a__and => true | id_s,id_a__U11 => true | id_s,id_U21 => true | id_s,id_a__U61 => true | id_s,id_a__U32 => true | id_s,id_U51 => true | id_s,id_isNatIListKind => true | id_s,id_0 => true | id_s,id_U12 => true | id_s,id_a__U52 => true | id_s,id_a__isNat => true | id_s,id_U43 => true | id_s,id_a__isNatIListKind => true | id_s,id_a__U12 => true | id_s,id_U31 => true | id_s,id_a__length => true | id_s,id_a__U42 => true | id_s,id_U53 => true | id_s,id_and => true | id_s,id_cons => true | id_s,id_U11 => true | id_s,id_a__U51 => true | id_s,id_a__U22 => true | id_s,id_U42 => true | id_s,id_length => true | id_s,id_tt => true | id_s,id_U22 => true | id_s,id_s => true | id_s,id_a__U41 => false | id_s,id_U52 => false | id_s,id_nil => false | id_s,id_zeros => false | id_s,id_isNatList => false | id_s,id_a__U53 => false | id_s,id_a__U31 => false | id_s,id_isNatIList => false | id_s,id_a__isNatKind => false | id_s,id_a__isNatList => false | id_s,id_U32 => false | id_s,id_mark => false | id_s,id_a__U43 => false | id_s,id_U61 => false | id_s,id_isNat => false | id_a__U41,id_a__zeros => true | id_a__U41,id_isNatKind => true | id_a__U41,id_a__isNatIList => true | id_a__U41,id_a__U21 => true | id_a__U41,id_U41 => true | id_a__U41,id_a__and => true | id_a__U41,id_a__U11 => true | id_a__U41,id_U21 => true | id_a__U41,id_a__U61 => true | id_a__U41,id_a__U32 => true | id_a__U41,id_U51 => true | id_a__U41,id_isNatIListKind => true | id_a__U41,id_0 => true | id_a__U41,id_U12 => true | id_a__U41,id_a__U52 => true | id_a__U41,id_a__isNat => true | id_a__U41,id_U43 => true | id_a__U41,id_a__isNatIListKind => true | id_a__U41,id_a__U12 => true | id_a__U41,id_U31 => true | id_a__U41,id_a__length => true | id_a__U41,id_a__U42 => true | id_a__U41,id_U53 => true | id_a__U41,id_and => true | id_a__U41,id_cons => true | id_a__U41,id_U11 => true | id_a__U41,id_a__U51 => true | id_a__U41,id_a__U22 => true | id_a__U41,id_U42 => true | id_a__U41,id_length => true | id_a__U41,id_tt => true | id_a__U41,id_U22 => true | id_a__U41,id_s => true | id_a__U41,id_a__U41 => true | id_a__U41,id_U52 => false | id_a__U41,id_nil => false | id_a__U41,id_zeros => false | id_a__U41,id_isNatList => false | id_a__U41,id_a__U53 => false | id_a__U41,id_a__U31 => false | id_a__U41,id_isNatIList => false | id_a__U41,id_a__isNatKind => false | id_a__U41,id_a__isNatList => false | id_a__U41,id_U32 => false | id_a__U41,id_mark => false | id_a__U41,id_a__U43 => false | id_a__U41,id_U61 => false | id_a__U41,id_isNat => false | id_U52,id_a__zeros => true | id_U52,id_isNatKind => true | id_U52,id_a__isNatIList => true | id_U52,id_a__U21 => true | id_U52,id_U41 => true | id_U52,id_a__and => true | id_U52,id_a__U11 => true | id_U52,id_U21 => true | id_U52,id_a__U61 => true | id_U52,id_a__U32 => true | id_U52,id_U51 => true | id_U52,id_isNatIListKind => true | id_U52,id_0 => true | id_U52,id_U12 => true | id_U52,id_a__U52 => true | id_U52,id_a__isNat => true | id_U52,id_U43 => true | id_U52,id_a__isNatIListKind => true | id_U52,id_a__U12 => true | id_U52,id_U31 => true | id_U52,id_a__length => true | id_U52,id_a__U42 => true | id_U52,id_U53 => true | id_U52,id_and => true | id_U52,id_cons => true | id_U52,id_U11 => true | id_U52,id_a__U51 => true | id_U52,id_a__U22 => true | id_U52,id_U42 => true | id_U52,id_length => true | id_U52,id_tt => true | id_U52,id_U22 => true | id_U52,id_s => true | id_U52,id_a__U41 => true | id_U52,id_U52 => true | id_U52,id_nil => false | id_U52,id_zeros => false | id_U52,id_isNatList => false | id_U52,id_a__U53 => false | id_U52,id_a__U31 => false | id_U52,id_isNatIList => false | id_U52,id_a__isNatKind => false | id_U52,id_a__isNatList => false | id_U52,id_U32 => false | id_U52,id_mark => false | id_U52,id_a__U43 => false | id_U52,id_U61 => false | id_U52,id_isNat => false | id_nil,id_a__zeros => true | id_nil,id_isNatKind => true | id_nil,id_a__isNatIList => true | id_nil,id_a__U21 => true | id_nil,id_U41 => true | id_nil,id_a__and => true | id_nil,id_a__U11 => true | id_nil,id_U21 => true | id_nil,id_a__U61 => true | id_nil,id_a__U32 => true | id_nil,id_U51 => true | id_nil,id_isNatIListKind => true | id_nil,id_0 => true | id_nil,id_U12 => true | id_nil,id_a__U52 => true | id_nil,id_a__isNat => true | id_nil,id_U43 => true | id_nil,id_a__isNatIListKind => true | id_nil,id_a__U12 => true | id_nil,id_U31 => true | id_nil,id_a__length => true | id_nil,id_a__U42 => true | id_nil,id_U53 => true | id_nil,id_and => true | id_nil,id_cons => true | id_nil,id_U11 => true | id_nil,id_a__U51 => true | id_nil,id_a__U22 => true | id_nil,id_U42 => true | id_nil,id_length => true | id_nil,id_tt => true | id_nil,id_U22 => true | id_nil,id_s => true | id_nil,id_a__U41 => true | id_nil,id_U52 => true | id_nil,id_nil => true | id_nil,id_zeros => false | id_nil,id_isNatList => false | id_nil,id_a__U53 => false | id_nil,id_a__U31 => false | id_nil,id_isNatIList => false | id_nil,id_a__isNatKind => false | id_nil,id_a__isNatList => false | id_nil,id_U32 => false | id_nil,id_mark => false | id_nil,id_a__U43 => false | id_nil,id_U61 => false | id_nil,id_isNat => false | id_zeros,id_a__zeros => true | id_zeros,id_isNatKind => true | id_zeros,id_a__isNatIList => true | id_zeros,id_a__U21 => true | id_zeros,id_U41 => true | id_zeros,id_a__and => true | id_zeros,id_a__U11 => true | id_zeros,id_U21 => true | id_zeros,id_a__U61 => true | id_zeros,id_a__U32 => true | id_zeros,id_U51 => true | id_zeros,id_isNatIListKind => true | id_zeros,id_0 => true | id_zeros,id_U12 => true | id_zeros,id_a__U52 => true | id_zeros,id_a__isNat => true | id_zeros,id_U43 => true | id_zeros,id_a__isNatIListKind => true | id_zeros,id_a__U12 => true | id_zeros,id_U31 => true | id_zeros,id_a__length => true | id_zeros,id_a__U42 => true | id_zeros,id_U53 => true | id_zeros,id_and => true | id_zeros,id_cons => true | id_zeros,id_U11 => true | id_zeros,id_a__U51 => true | id_zeros,id_a__U22 => true | id_zeros,id_U42 => true | id_zeros,id_length => true | id_zeros,id_tt => true | id_zeros,id_U22 => true | id_zeros,id_s => true | id_zeros,id_a__U41 => true | id_zeros,id_U52 => true | id_zeros,id_nil => true | id_zeros,id_zeros => true | id_zeros,id_isNatList => false | id_zeros,id_a__U53 => false | id_zeros,id_a__U31 => false | id_zeros,id_isNatIList => false | id_zeros,id_a__isNatKind => false | id_zeros,id_a__isNatList => false | id_zeros,id_U32 => false | id_zeros,id_mark => false | id_zeros,id_a__U43 => false | id_zeros,id_U61 => false | id_zeros,id_isNat => false | id_isNatList,id_a__zeros => true | id_isNatList,id_isNatKind => true | id_isNatList,id_a__isNatIList => true | id_isNatList,id_a__U21 => true | id_isNatList,id_U41 => true | id_isNatList,id_a__and => true | id_isNatList,id_a__U11 => true | id_isNatList,id_U21 => true | id_isNatList,id_a__U61 => true | id_isNatList,id_a__U32 => true | id_isNatList,id_U51 => true | id_isNatList,id_isNatIListKind => true | id_isNatList,id_0 => true | id_isNatList,id_U12 => true | id_isNatList,id_a__U52 => true | id_isNatList,id_a__isNat => true | id_isNatList,id_U43 => true | id_isNatList,id_a__isNatIListKind => true | id_isNatList,id_a__U12 => true | id_isNatList,id_U31 => true | id_isNatList,id_a__length => true | id_isNatList,id_a__U42 => true | id_isNatList,id_U53 => true | id_isNatList,id_and => true | id_isNatList,id_cons => true | id_isNatList,id_U11 => true | id_isNatList,id_a__U51 => true | id_isNatList,id_a__U22 => true | id_isNatList,id_U42 => true | id_isNatList,id_length => true | id_isNatList,id_tt => true | id_isNatList,id_U22 => true | id_isNatList,id_s => true | id_isNatList,id_a__U41 => true | id_isNatList,id_U52 => true | id_isNatList,id_nil => true | id_isNatList,id_zeros => true | id_isNatList,id_isNatList => true | id_isNatList,id_a__U53 => false | id_isNatList,id_a__U31 => false | id_isNatList,id_isNatIList => false | id_isNatList,id_a__isNatKind => false | id_isNatList,id_a__isNatList => false | id_isNatList,id_U32 => false | id_isNatList,id_mark => false | id_isNatList,id_a__U43 => false | id_isNatList,id_U61 => false | id_isNatList,id_isNat => false | id_a__U53,id_a__zeros => true | id_a__U53,id_isNatKind => true | id_a__U53,id_a__isNatIList => true | id_a__U53,id_a__U21 => true | id_a__U53,id_U41 => true | id_a__U53,id_a__and => true | id_a__U53,id_a__U11 => true | id_a__U53,id_U21 => true | id_a__U53,id_a__U61 => true | id_a__U53,id_a__U32 => true | id_a__U53,id_U51 => true | id_a__U53,id_isNatIListKind => true | id_a__U53,id_0 => true | id_a__U53,id_U12 => true | id_a__U53,id_a__U52 => true | id_a__U53,id_a__isNat => true | id_a__U53,id_U43 => true | id_a__U53,id_a__isNatIListKind => true | id_a__U53,id_a__U12 => true | id_a__U53,id_U31 => true | id_a__U53,id_a__length => true | id_a__U53,id_a__U42 => true | id_a__U53,id_U53 => true | id_a__U53,id_and => true | id_a__U53,id_cons => true | id_a__U53,id_U11 => true | id_a__U53,id_a__U51 => true | id_a__U53,id_a__U22 => true | id_a__U53,id_U42 => true | id_a__U53,id_length => true | id_a__U53,id_tt => true | id_a__U53,id_U22 => true | id_a__U53,id_s => true | id_a__U53,id_a__U41 => true | id_a__U53,id_U52 => true | id_a__U53,id_nil => true | id_a__U53,id_zeros => true | id_a__U53,id_isNatList => true | id_a__U53,id_a__U53 => true | id_a__U53,id_a__U31 => false | id_a__U53,id_isNatIList => false | id_a__U53,id_a__isNatKind => false | id_a__U53,id_a__isNatList => false | id_a__U53,id_U32 => false | id_a__U53,id_mark => false | id_a__U53,id_a__U43 => false | id_a__U53,id_U61 => false | id_a__U53,id_isNat => false | id_a__U31,id_a__zeros => true | id_a__U31,id_isNatKind => true | id_a__U31,id_a__isNatIList => true | id_a__U31,id_a__U21 => true | id_a__U31,id_U41 => true | id_a__U31,id_a__and => true | id_a__U31,id_a__U11 => true | id_a__U31,id_U21 => true | id_a__U31,id_a__U61 => true | id_a__U31,id_a__U32 => true | id_a__U31,id_U51 => true | id_a__U31,id_isNatIListKind => true | id_a__U31,id_0 => true | id_a__U31,id_U12 => true | id_a__U31,id_a__U52 => true | id_a__U31,id_a__isNat => true | id_a__U31,id_U43 => true | id_a__U31,id_a__isNatIListKind => true | id_a__U31,id_a__U12 => true | id_a__U31,id_U31 => true | id_a__U31,id_a__length => true | id_a__U31,id_a__U42 => true | id_a__U31,id_U53 => true | id_a__U31,id_and => true | id_a__U31,id_cons => true | id_a__U31,id_U11 => true | id_a__U31,id_a__U51 => true | id_a__U31,id_a__U22 => true | id_a__U31,id_U42 => true | id_a__U31,id_length => true | id_a__U31,id_tt => true | id_a__U31,id_U22 => true | id_a__U31,id_s => true | id_a__U31,id_a__U41 => true | id_a__U31,id_U52 => true | id_a__U31,id_nil => true | id_a__U31,id_zeros => true | id_a__U31,id_isNatList => true | id_a__U31,id_a__U53 => true | id_a__U31,id_a__U31 => true | id_a__U31,id_isNatIList => false | id_a__U31,id_a__isNatKind => false | id_a__U31,id_a__isNatList => false | id_a__U31,id_U32 => false | id_a__U31,id_mark => false | id_a__U31,id_a__U43 => false | id_a__U31,id_U61 => false | id_a__U31,id_isNat => false | id_isNatIList,id_a__zeros => true | id_isNatIList,id_isNatKind => true | id_isNatIList,id_a__isNatIList => true | id_isNatIList,id_a__U21 => true | id_isNatIList,id_U41 => true | id_isNatIList,id_a__and => true | id_isNatIList,id_a__U11 => true | id_isNatIList,id_U21 => true | id_isNatIList,id_a__U61 => true | id_isNatIList,id_a__U32 => true | id_isNatIList,id_U51 => true | id_isNatIList,id_isNatIListKind => true | id_isNatIList,id_0 => true | id_isNatIList,id_U12 => true | id_isNatIList,id_a__U52 => true | id_isNatIList,id_a__isNat => true | id_isNatIList,id_U43 => true | id_isNatIList,id_a__isNatIListKind => true | id_isNatIList,id_a__U12 => true | id_isNatIList,id_U31 => true | id_isNatIList,id_a__length => true | id_isNatIList,id_a__U42 => true | id_isNatIList,id_U53 => true | id_isNatIList,id_and => true | id_isNatIList,id_cons => true | id_isNatIList,id_U11 => true | id_isNatIList,id_a__U51 => true | id_isNatIList,id_a__U22 => true | id_isNatIList,id_U42 => true | id_isNatIList,id_length => true | id_isNatIList,id_tt => true | id_isNatIList,id_U22 => true | id_isNatIList,id_s => true | id_isNatIList,id_a__U41 => true | id_isNatIList,id_U52 => true | id_isNatIList,id_nil => true | id_isNatIList,id_zeros => true | id_isNatIList,id_isNatList => true | id_isNatIList,id_a__U53 => true | id_isNatIList,id_a__U31 => true | id_isNatIList,id_isNatIList => true | id_isNatIList,id_a__isNatKind => false | id_isNatIList,id_a__isNatList => false | id_isNatIList,id_U32 => false | id_isNatIList,id_mark => false | id_isNatIList,id_a__U43 => false | id_isNatIList,id_U61 => false | id_isNatIList,id_isNat => false | id_a__isNatKind,id_a__zeros => true | id_a__isNatKind,id_isNatKind => true | id_a__isNatKind,id_a__isNatIList => true | id_a__isNatKind,id_a__U21 => true | id_a__isNatKind,id_U41 => true | id_a__isNatKind,id_a__and => true | id_a__isNatKind,id_a__U11 => true | id_a__isNatKind,id_U21 => true | id_a__isNatKind,id_a__U61 => true | id_a__isNatKind,id_a__U32 => true | id_a__isNatKind,id_U51 => true | id_a__isNatKind,id_isNatIListKind => true | id_a__isNatKind,id_0 => true | id_a__isNatKind,id_U12 => true | id_a__isNatKind,id_a__U52 => true | id_a__isNatKind,id_a__isNat => true | id_a__isNatKind,id_U43 => true | id_a__isNatKind,id_a__isNatIListKind => true | id_a__isNatKind,id_a__U12 => true | id_a__isNatKind,id_U31 => true | id_a__isNatKind,id_a__length => true | id_a__isNatKind,id_a__U42 => true | id_a__isNatKind,id_U53 => true | id_a__isNatKind,id_and => true | id_a__isNatKind,id_cons => true | id_a__isNatKind,id_U11 => true | id_a__isNatKind,id_a__U51 => true | id_a__isNatKind,id_a__U22 => true | id_a__isNatKind,id_U42 => true | id_a__isNatKind,id_length => true | id_a__isNatKind,id_tt => true | id_a__isNatKind,id_U22 => true | id_a__isNatKind,id_s => true | id_a__isNatKind,id_a__U41 => true | id_a__isNatKind,id_U52 => true | id_a__isNatKind,id_nil => true | id_a__isNatKind,id_zeros => true | id_a__isNatKind,id_isNatList => true | id_a__isNatKind,id_a__U53 => true | id_a__isNatKind,id_a__U31 => true | id_a__isNatKind,id_isNatIList => true | id_a__isNatKind,id_a__isNatKind => true | id_a__isNatKind,id_a__isNatList => false | id_a__isNatKind,id_U32 => false | id_a__isNatKind,id_mark => false | id_a__isNatKind,id_a__U43 => false | id_a__isNatKind,id_U61 => false | id_a__isNatKind,id_isNat => false | id_a__isNatList,id_a__zeros => true | id_a__isNatList,id_isNatKind => true | id_a__isNatList,id_a__isNatIList => true | id_a__isNatList,id_a__U21 => true | id_a__isNatList,id_U41 => true | id_a__isNatList,id_a__and => true | id_a__isNatList,id_a__U11 => true | id_a__isNatList,id_U21 => true | id_a__isNatList,id_a__U61 => true | id_a__isNatList,id_a__U32 => true | id_a__isNatList,id_U51 => true | id_a__isNatList,id_isNatIListKind => true | id_a__isNatList,id_0 => true | id_a__isNatList,id_U12 => true | id_a__isNatList,id_a__U52 => true | id_a__isNatList,id_a__isNat => true | id_a__isNatList,id_U43 => true | id_a__isNatList,id_a__isNatIListKind => true | id_a__isNatList,id_a__U12 => true | id_a__isNatList,id_U31 => true | id_a__isNatList,id_a__length => true | id_a__isNatList,id_a__U42 => true | id_a__isNatList,id_U53 => true | id_a__isNatList,id_and => true | id_a__isNatList,id_cons => true | id_a__isNatList,id_U11 => true | id_a__isNatList,id_a__U51 => true | id_a__isNatList,id_a__U22 => true | id_a__isNatList,id_U42 => true | id_a__isNatList,id_length => true | id_a__isNatList,id_tt => true | id_a__isNatList,id_U22 => true | id_a__isNatList,id_s => true | id_a__isNatList,id_a__U41 => true | id_a__isNatList,id_U52 => true | id_a__isNatList,id_nil => true | id_a__isNatList,id_zeros => true | id_a__isNatList,id_isNatList => true | id_a__isNatList,id_a__U53 => true | id_a__isNatList,id_a__U31 => true | id_a__isNatList,id_isNatIList => true | id_a__isNatList,id_a__isNatKind => true | id_a__isNatList,id_a__isNatList => true | id_a__isNatList,id_U32 => false | id_a__isNatList,id_mark => false | id_a__isNatList,id_a__U43 => false | id_a__isNatList,id_U61 => false | id_a__isNatList,id_isNat => false | id_U32,id_a__zeros => true | id_U32,id_isNatKind => true | id_U32,id_a__isNatIList => true | id_U32,id_a__U21 => true | id_U32,id_U41 => true | id_U32,id_a__and => true | id_U32,id_a__U11 => true | id_U32,id_U21 => true | id_U32,id_a__U61 => true | id_U32,id_a__U32 => true | id_U32,id_U51 => true | id_U32,id_isNatIListKind => true | id_U32,id_0 => true | id_U32,id_U12 => true | id_U32,id_a__U52 => true | id_U32,id_a__isNat => true | id_U32,id_U43 => true | id_U32,id_a__isNatIListKind => true | id_U32,id_a__U12 => true | id_U32,id_U31 => true | id_U32,id_a__length => true | id_U32,id_a__U42 => true | id_U32,id_U53 => true | id_U32,id_and => true | id_U32,id_cons => true | id_U32,id_U11 => true | id_U32,id_a__U51 => true | id_U32,id_a__U22 => true | id_U32,id_U42 => true | id_U32,id_length => true | id_U32,id_tt => true | id_U32,id_U22 => true | id_U32,id_s => true | id_U32,id_a__U41 => true | id_U32,id_U52 => true | id_U32,id_nil => true | id_U32,id_zeros => true | id_U32,id_isNatList => true | id_U32,id_a__U53 => true | id_U32,id_a__U31 => true | id_U32,id_isNatIList => true | id_U32,id_a__isNatKind => true | id_U32,id_a__isNatList => true | id_U32,id_U32 => true | id_U32,id_mark => false | id_U32,id_a__U43 => false | id_U32,id_U61 => false | id_U32,id_isNat => false | id_mark,id_a__zeros => true | id_mark,id_isNatKind => true | id_mark,id_a__isNatIList => true | id_mark,id_a__U21 => true | id_mark,id_U41 => true | id_mark,id_a__and => true | id_mark,id_a__U11 => true | id_mark,id_U21 => true | id_mark,id_a__U61 => true | id_mark,id_a__U32 => true | id_mark,id_U51 => true | id_mark,id_isNatIListKind => true | id_mark,id_0 => true | id_mark,id_U12 => true | id_mark,id_a__U52 => true | id_mark,id_a__isNat => true | id_mark,id_U43 => true | id_mark,id_a__isNatIListKind => true | id_mark,id_a__U12 => true | id_mark,id_U31 => true | id_mark,id_a__length => true | id_mark,id_a__U42 => true | id_mark,id_U53 => true | id_mark,id_and => true | id_mark,id_cons => true | id_mark,id_U11 => true | id_mark,id_a__U51 => true | id_mark,id_a__U22 => true | id_mark,id_U42 => true | id_mark,id_length => true | id_mark,id_tt => true | id_mark,id_U22 => true | id_mark,id_s => true | id_mark,id_a__U41 => true | id_mark,id_U52 => true | id_mark,id_nil => true | id_mark,id_zeros => true | id_mark,id_isNatList => true | id_mark,id_a__U53 => true | id_mark,id_a__U31 => true | id_mark,id_isNatIList => true | id_mark,id_a__isNatKind => true | id_mark,id_a__isNatList => true | id_mark,id_U32 => true | id_mark,id_mark => true | id_mark,id_a__U43 => false | id_mark,id_U61 => false | id_mark,id_isNat => false | id_a__U43,id_a__zeros => true | id_a__U43,id_isNatKind => true | id_a__U43,id_a__isNatIList => true | id_a__U43,id_a__U21 => true | id_a__U43,id_U41 => true | id_a__U43,id_a__and => true | id_a__U43,id_a__U11 => true | id_a__U43,id_U21 => true | id_a__U43,id_a__U61 => true | id_a__U43,id_a__U32 => true | id_a__U43,id_U51 => true | id_a__U43,id_isNatIListKind => true | id_a__U43,id_0 => true | id_a__U43,id_U12 => true | id_a__U43,id_a__U52 => true | id_a__U43,id_a__isNat => true | id_a__U43,id_U43 => true | id_a__U43,id_a__isNatIListKind => true | id_a__U43,id_a__U12 => true | id_a__U43,id_U31 => true | id_a__U43,id_a__length => true | id_a__U43,id_a__U42 => true | id_a__U43,id_U53 => true | id_a__U43,id_and => true | id_a__U43,id_cons => true | id_a__U43,id_U11 => true | id_a__U43,id_a__U51 => true | id_a__U43,id_a__U22 => true | id_a__U43,id_U42 => true | id_a__U43,id_length => true | id_a__U43,id_tt => true | id_a__U43,id_U22 => true | id_a__U43,id_s => true | id_a__U43,id_a__U41 => true | id_a__U43,id_U52 => true | id_a__U43,id_nil => true | id_a__U43,id_zeros => true | id_a__U43,id_isNatList => true | id_a__U43,id_a__U53 => true | id_a__U43,id_a__U31 => true | id_a__U43,id_isNatIList => true | id_a__U43,id_a__isNatKind => true | id_a__U43,id_a__isNatList => true | id_a__U43,id_U32 => true | id_a__U43,id_mark => true | id_a__U43,id_a__U43 => true | id_a__U43,id_U61 => false | id_a__U43,id_isNat => false | id_U61,id_a__zeros => true | id_U61,id_isNatKind => true | id_U61,id_a__isNatIList => true | id_U61,id_a__U21 => true | id_U61,id_U41 => true | id_U61,id_a__and => true | id_U61,id_a__U11 => true | id_U61,id_U21 => true | id_U61,id_a__U61 => true | id_U61,id_a__U32 => true | id_U61,id_U51 => true | id_U61,id_isNatIListKind => true | id_U61,id_0 => true | id_U61,id_U12 => true | id_U61,id_a__U52 => true | id_U61,id_a__isNat => true | id_U61,id_U43 => true | id_U61,id_a__isNatIListKind => true | id_U61,id_a__U12 => true | id_U61,id_U31 => true | id_U61,id_a__length => true | id_U61,id_a__U42 => true | id_U61,id_U53 => true | id_U61,id_and => true | id_U61,id_cons => true | id_U61,id_U11 => true | id_U61,id_a__U51 => true | id_U61,id_a__U22 => true | id_U61,id_U42 => true | id_U61,id_length => true | id_U61,id_tt => true | id_U61,id_U22 => true | id_U61,id_s => true | id_U61,id_a__U41 => true | id_U61,id_U52 => true | id_U61,id_nil => true | id_U61,id_zeros => true | id_U61,id_isNatList => true | id_U61,id_a__U53 => true | id_U61,id_a__U31 => true | id_U61,id_isNatIList => true | id_U61,id_a__isNatKind => true | id_U61,id_a__isNatList => true | id_U61,id_U32 => true | id_U61,id_mark => true | id_U61,id_a__U43 => true | id_U61,id_U61 => true | id_U61,id_isNat => false | id_isNat,id_a__zeros => true | id_isNat,id_isNatKind => true | id_isNat,id_a__isNatIList => true | id_isNat,id_a__U21 => true | id_isNat,id_U41 => true | id_isNat,id_a__and => true | id_isNat,id_a__U11 => true | id_isNat,id_U21 => true | id_isNat,id_a__U61 => true | id_isNat,id_a__U32 => true | id_isNat,id_U51 => true | id_isNat,id_isNatIListKind => true | id_isNat,id_0 => true | id_isNat,id_U12 => true | id_isNat,id_a__U52 => true | id_isNat,id_a__isNat => true | id_isNat,id_U43 => true | id_isNat,id_a__isNatIListKind => true | id_isNat,id_a__U12 => true | id_isNat,id_U31 => true | id_isNat,id_a__length => true | id_isNat,id_a__U42 => true | id_isNat,id_U53 => true | id_isNat,id_and => true | id_isNat,id_cons => true | id_isNat,id_U11 => true | id_isNat,id_a__U51 => true | id_isNat,id_a__U22 => true | id_isNat,id_U42 => true | id_isNat,id_length => true | id_isNat,id_tt => true | id_isNat,id_U22 => true | id_isNat,id_s => true | id_isNat,id_a__U41 => true | id_isNat,id_U52 => true | id_isNat,id_nil => true | id_isNat,id_zeros => true | id_isNat,id_isNatList => true | id_isNat,id_a__U53 => true | id_isNat,id_a__U31 => true | id_isNat,id_isNatIList => true | id_isNat,id_a__isNatKind => true | id_isNat,id_a__isNatList => true | id_isNat,id_U32 => true | id_isNat,id_mark => true | id_isNat,id_a__U43 => true | id_isNat,id_U61 => true | id_isNat,id_isNat => 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 := (* a__zeros -> cons(0,zeros) *) | R_xml_0_rule_0 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_0 nil)::(algebra.Alg.Term algebra.F.id_zeros nil)::nil)) (algebra.Alg.Term algebra.F.id_a__zeros nil) (* a__U11(tt,V1_) -> a__U12(a__isNatList(V1_)) *) | R_xml_0_rule_1 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_a__isNatList ((algebra.Alg.Var 1)::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 1)::nil)) (* a__U12(tt) -> tt *) | R_xml_0_rule_2 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)) (* a__U21(tt,V1_) -> a__U22(a__isNat(V1_)) *) | R_xml_0_rule_3 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_a__isNat ((algebra.Alg.Var 1)::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 1)::nil)) (* a__U22(tt) -> tt *) | R_xml_0_rule_4 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)) (* a__U31(tt,V_) -> a__U32(a__isNatList(V_)) *) | R_xml_0_rule_5 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_a__isNatList ((algebra.Alg.Var 2)::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 2)::nil)) (* a__U32(tt) -> tt *) | R_xml_0_rule_6 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)) (* a__U41(tt,V1_,V2_) -> a__U42(a__isNat(V1_),V2_) *) | R_xml_0_rule_7 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isNat ((algebra.Alg.Var 1)::nil)):: (algebra.Alg.Var 3)::nil)) (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 1)::(algebra.Alg.Var 3)::nil)) (* a__U42(tt,V2_) -> a__U43(a__isNatIList(V2_)) *) | R_xml_0_rule_8 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_a__isNatIList ((algebra.Alg.Var 3)::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 3)::nil)) (* a__U43(tt) -> tt *) | R_xml_0_rule_9 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)) (* a__U51(tt,V1_,V2_) -> a__U52(a__isNat(V1_),V2_) *) | R_xml_0_rule_10 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNat ((algebra.Alg.Var 1)::nil)):: (algebra.Alg.Var 3)::nil)) (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 1)::(algebra.Alg.Var 3)::nil)) (* a__U52(tt,V2_) -> a__U53(a__isNatList(V2_)) *) | R_xml_0_rule_11 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_a__isNatList ((algebra.Alg.Var 3)::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 3)::nil)) (* a__U53(tt) -> tt *) | R_xml_0_rule_12 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)) (* a__U61(tt,L_) -> s(a__length(mark(L_))) *) | R_xml_0_rule_13 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 4)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 4)::nil)) (* a__and(tt,X_) -> mark(X_) *) | R_xml_0_rule_14 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 5)::nil)) (* a__isNat(0) -> tt *) | R_xml_0_rule_15 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isNat ((algebra.Alg.Term algebra.F.id_0 nil)::nil)) (* a__isNat(length(V1_)) -> a__U11(a__isNatIListKind(V1_),V1_) *) | R_xml_0_rule_16 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__isNat ((algebra.Alg.Term algebra.F.id_length ((algebra.Alg.Var 1)::nil))::nil)) (* a__isNat(s(V1_)) -> a__U21(a__isNatKind(V1_),V1_) *) | R_xml_0_rule_17 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Var 1)::nil)):: (algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__isNat ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Var 1)::nil))::nil)) (* a__isNatIList(V_) -> a__U31(a__isNatIListKind(V_),V_) *) | R_xml_0_rule_18 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind ((algebra.Alg.Var 2)::nil))::(algebra.Alg.Var 2)::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList ((algebra.Alg.Var 2)::nil)) (* a__isNatIList(zeros) -> tt *) | R_xml_0_rule_19 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isNatIList ((algebra.Alg.Term algebra.F.id_zeros nil)::nil)) (* a__isNatIList(cons(V1_,V2_)) -> a__U41(a__and(a__isNatKind(V1_),isNatIListKind(V2_)),V1_,V2_) *) | R_xml_0_rule_20 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Var 1)::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind ((algebra.Alg.Var 3)::nil))::nil))::(algebra.Alg.Var 1):: (algebra.Alg.Var 3)::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 1):: (algebra.Alg.Var 3)::nil))::nil)) (* a__isNatIListKind(nil) -> tt *) | R_xml_0_rule_21 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isNatIListKind ((algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* a__isNatIListKind(zeros) -> tt *) | R_xml_0_rule_22 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isNatIListKind ((algebra.Alg.Term algebra.F.id_zeros nil)::nil)) (* a__isNatIListKind(cons(V1_,V2_)) -> a__and(a__isNatKind(V1_),isNatIListKind(V2_)) *) | R_xml_0_rule_23 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Var 1)::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind ((algebra.Alg.Var 3)::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 1):: (algebra.Alg.Var 3)::nil))::nil)) (* a__isNatKind(0) -> tt *) | R_xml_0_rule_24 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Term algebra.F.id_0 nil)::nil)) (* a__isNatKind(length(V1_)) -> a__isNatIListKind(V1_) *) | R_xml_0_rule_25 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__isNatIListKind ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Term algebra.F.id_length ((algebra.Alg.Var 1)::nil))::nil)) (* a__isNatKind(s(V1_)) -> a__isNatKind(V1_) *) | R_xml_0_rule_26 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Var 1)::nil))::nil)) (* a__isNatList(nil) -> tt *) | R_xml_0_rule_27 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isNatList ((algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* a__isNatList(cons(V1_,V2_)) -> a__U51(a__and(a__isNatKind(V1_),isNatIListKind(V2_)),V1_,V2_) *) | R_xml_0_rule_28 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Var 1)::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind ((algebra.Alg.Var 3)::nil))::nil))::(algebra.Alg.Var 1):: (algebra.Alg.Var 3)::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 1):: (algebra.Alg.Var 3)::nil))::nil)) (* a__length(nil) -> 0 *) | R_xml_0_rule_29 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_0 nil) (algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* a__length(cons(N_,L_)) -> a__U61(a__and(a__and(a__isNatList(L_),isNatIListKind(L_)), and(isNat(N_),isNatKind(N_))),L_) *) | R_xml_0_rule_30 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatList ((algebra.Alg.Var 4)::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind ((algebra.Alg.Var 4)::nil))::nil))::(algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isNat ((algebra.Alg.Var 6)::nil))::(algebra.Alg.Term algebra.F.id_isNatKind ((algebra.Alg.Var 6)::nil))::nil))::nil)):: (algebra.Alg.Var 4)::nil)) (algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 6):: (algebra.Alg.Var 4)::nil))::nil)) (* mark(zeros) -> a__zeros *) | R_xml_0_rule_31 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__zeros nil) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_zeros nil)::nil)) (* mark(U11(X1_,X2_)) -> a__U11(mark(X1_),X2_) *) | R_xml_0_rule_32 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8)::nil))::nil)) (* mark(U12(X_)) -> a__U12(mark(X_)) *) | R_xml_0_rule_33 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U12 ((algebra.Alg.Var 5)::nil))::nil)) (* mark(isNatList(X_)) -> a__isNatList(X_) *) | R_xml_0_rule_34 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__isNatList ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isNatList ((algebra.Alg.Var 5)::nil))::nil)) (* mark(U21(X1_,X2_)) -> a__U21(mark(X1_),X2_) *) | R_xml_0_rule_35 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8)::nil))::nil)) (* mark(U22(X_)) -> a__U22(mark(X_)) *) | R_xml_0_rule_36 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Var 5)::nil))::nil)) (* mark(isNat(X_)) -> a__isNat(X_) *) | R_xml_0_rule_37 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__isNat ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isNat ((algebra.Alg.Var 5)::nil))::nil)) (* mark(U31(X1_,X2_)) -> a__U31(mark(X1_),X2_) *) | R_xml_0_rule_38 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8)::nil))::nil)) (* mark(U32(X_)) -> a__U32(mark(X_)) *) | R_xml_0_rule_39 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U32 ((algebra.Alg.Var 5)::nil))::nil)) (* mark(U41(X1_,X2_,X3_)) -> a__U41(mark(X1_),X2_,X3_) *) | R_xml_0_rule_40 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8):: (algebra.Alg.Var 9)::nil))::nil)) (* mark(U42(X1_,X2_)) -> a__U42(mark(X1_),X2_) *) | R_xml_0_rule_41 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8)::nil))::nil)) (* mark(U43(X_)) -> a__U43(mark(X_)) *) | R_xml_0_rule_42 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U43 ((algebra.Alg.Var 5)::nil))::nil)) (* mark(isNatIList(X_)) -> a__isNatIList(X_) *) | R_xml_0_rule_43 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__isNatIList ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isNatIList ((algebra.Alg.Var 5)::nil))::nil)) (* mark(U51(X1_,X2_,X3_)) -> a__U51(mark(X1_),X2_,X3_) *) | R_xml_0_rule_44 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8):: (algebra.Alg.Var 9)::nil))::nil)) (* mark(U52(X1_,X2_)) -> a__U52(mark(X1_),X2_) *) | R_xml_0_rule_45 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8)::nil))::nil)) (* mark(U53(X_)) -> a__U53(mark(X_)) *) | R_xml_0_rule_46 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U53 ((algebra.Alg.Var 5)::nil))::nil)) (* mark(U61(X1_,X2_)) -> a__U61(mark(X1_),X2_) *) | R_xml_0_rule_47 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8)::nil))::nil)) (* mark(length(X_)) -> a__length(mark(X_)) *) | R_xml_0_rule_48 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_length ((algebra.Alg.Var 5)::nil))::nil)) (* mark(and(X1_,X2_)) -> a__and(mark(X1_),X2_) *) | R_xml_0_rule_49 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Var 7)::(algebra.Alg.Var 8)::nil))::nil)) (* mark(isNatIListKind(X_)) -> a__isNatIListKind(X_) *) | R_xml_0_rule_50 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__isNatIListKind ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isNatIListKind ((algebra.Alg.Var 5)::nil))::nil)) (* mark(isNatKind(X_)) -> a__isNatKind(X_) *) | R_xml_0_rule_51 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isNatKind ((algebra.Alg.Var 5)::nil))::nil)) (* mark(cons(X1_,X2_)) -> cons(mark(X1_),X2_) *) | R_xml_0_rule_52 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil))::nil)) (* mark(0) -> 0 *) | R_xml_0_rule_53 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_0 nil) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_0 nil)::nil)) (* mark(tt) -> tt *) | R_xml_0_rule_54 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt nil)::nil)) (* mark(s(X_)) -> s(mark(X_)) *) | R_xml_0_rule_55 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Var 5)::nil))::nil)) (* mark(nil) -> nil *) | R_xml_0_rule_56 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_nil nil) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* a__zeros -> zeros *) | R_xml_0_rule_57 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_zeros nil) (algebra.Alg.Term algebra.F.id_a__zeros nil) (* a__U11(X1_,X2_) -> U11(X1_,X2_) *) | R_xml_0_rule_58 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)) (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)) (* a__U12(X_) -> U12(X_) *) | R_xml_0_rule_59 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U12 ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Var 5)::nil)) (* a__isNatList(X_) -> isNatList(X_) *) | R_xml_0_rule_60 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_isNatList ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList ((algebra.Alg.Var 5)::nil)) (* a__U21(X1_,X2_) -> U21(X1_,X2_) *) | R_xml_0_rule_61 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)) (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)) (* a__U22(X_) -> U22(X_) *) | R_xml_0_rule_62 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Var 5)::nil)) (* a__isNat(X_) -> isNat(X_) *) | R_xml_0_rule_63 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_isNat ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_a__isNat ((algebra.Alg.Var 5)::nil)) (* a__U31(X1_,X2_) -> U31(X1_,X2_) *) | R_xml_0_rule_64 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)) (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)) (* a__U32(X_) -> U32(X_) *) | R_xml_0_rule_65 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U32 ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Var 5)::nil)) (* a__U41(X1_,X2_,X3_) -> U41(X1_,X2_,X3_) *) | R_xml_0_rule_66 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)) (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)) (* a__U42(X1_,X2_) -> U42(X1_,X2_) *) | R_xml_0_rule_67 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)) (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)) (* a__U43(X_) -> U43(X_) *) | R_xml_0_rule_68 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U43 ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Var 5)::nil)) (* a__isNatIList(X_) -> isNatIList(X_) *) | R_xml_0_rule_69 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_isNatIList ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList ((algebra.Alg.Var 5)::nil)) (* a__U51(X1_,X2_,X3_) -> U51(X1_,X2_,X3_) *) | R_xml_0_rule_70 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)) (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)) (* a__U52(X1_,X2_) -> U52(X1_,X2_) *) | R_xml_0_rule_71 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)) (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)) (* a__U53(X_) -> U53(X_) *) | R_xml_0_rule_72 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U53 ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Var 5)::nil)) (* a__U61(X1_,X2_) -> U61(X1_,X2_) *) | R_xml_0_rule_73 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)) (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)) (* a__length(X_) -> length(X_) *) | R_xml_0_rule_74 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_length ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Var 5)::nil)) (* a__and(X1_,X2_) -> and(X1_,X2_) *) | R_xml_0_rule_75 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)) (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)) (* a__isNatIListKind(X_) -> isNatIListKind(X_) *) | R_xml_0_rule_76 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_isNatIListKind ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind ((algebra.Alg.Var 5)::nil)) (* a__isNatKind(X_) -> isNatKind(X_) *) | R_xml_0_rule_77 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_isNatKind ((algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Var 5)::nil)) . Definition R_xml_0_rule_as_list_0 := ((algebra.Alg.Term algebra.F.id_a__zeros nil), (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_0 nil)::(algebra.Alg.Term algebra.F.id_zeros nil)::nil)))::nil. Definition R_xml_0_rule_as_list_1 := ((algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_a__isNatList ((algebra.Alg.Var 1)::nil))::nil))):: R_xml_0_rule_as_list_0. Definition R_xml_0_rule_as_list_2 := ((algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_1. Definition R_xml_0_rule_as_list_3 := ((algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_a__isNat ((algebra.Alg.Var 1)::nil))::nil))):: R_xml_0_rule_as_list_2. Definition R_xml_0_rule_as_list_4 := ((algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_3. Definition R_xml_0_rule_as_list_5 := ((algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 2)::nil)), (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_a__isNatList ((algebra.Alg.Var 2)::nil))::nil))):: R_xml_0_rule_as_list_4. Definition R_xml_0_rule_as_list_6 := ((algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_5. Definition R_xml_0_rule_as_list_7 := ((algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 1)::(algebra.Alg.Var 3)::nil)), (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isNat ((algebra.Alg.Var 1)::nil)):: (algebra.Alg.Var 3)::nil)))::R_xml_0_rule_as_list_6. Definition R_xml_0_rule_as_list_8 := ((algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 3)::nil)), (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_a__isNatIList ((algebra.Alg.Var 3)::nil))::nil))):: R_xml_0_rule_as_list_7. Definition R_xml_0_rule_as_list_9 := ((algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_8. Definition R_xml_0_rule_as_list_10 := ((algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 1)::(algebra.Alg.Var 3)::nil)), (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNat ((algebra.Alg.Var 1)::nil)):: (algebra.Alg.Var 3)::nil)))::R_xml_0_rule_as_list_9. Definition R_xml_0_rule_as_list_11 := ((algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 3)::nil)), (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_a__isNatList ((algebra.Alg.Var 3)::nil))::nil))):: R_xml_0_rule_as_list_10. Definition R_xml_0_rule_as_list_12 := ((algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_11. Definition R_xml_0_rule_as_list_13 := ((algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 4)::nil)), (algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 4)::nil))::nil))::nil)))::R_xml_0_rule_as_list_12. Definition R_xml_0_rule_as_list_14 := ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 5)::nil)), (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 5)::nil))):: R_xml_0_rule_as_list_13. Definition R_xml_0_rule_as_list_15 := ((algebra.Alg.Term algebra.F.id_a__isNat ((algebra.Alg.Term algebra.F.id_0 nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_14. Definition R_xml_0_rule_as_list_16 := ((algebra.Alg.Term algebra.F.id_a__isNat ((algebra.Alg.Term algebra.F.id_length ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind ((algebra.Alg.Var 1)::nil)):: (algebra.Alg.Var 1)::nil)))::R_xml_0_rule_as_list_15. Definition R_xml_0_rule_as_list_17 := ((algebra.Alg.Term algebra.F.id_a__isNat ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Var 1)::nil)):: (algebra.Alg.Var 1)::nil)))::R_xml_0_rule_as_list_16. Definition R_xml_0_rule_as_list_18 := ((algebra.Alg.Term algebra.F.id_a__isNatIList ((algebra.Alg.Var 2)::nil)), (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind ((algebra.Alg.Var 2)::nil)):: (algebra.Alg.Var 2)::nil)))::R_xml_0_rule_as_list_17. Definition R_xml_0_rule_as_list_19 := ((algebra.Alg.Term algebra.F.id_a__isNatIList ((algebra.Alg.Term algebra.F.id_zeros nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_18. Definition R_xml_0_rule_as_list_20 := ((algebra.Alg.Term algebra.F.id_a__isNatIList ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 1):: (algebra.Alg.Var 3)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind ((algebra.Alg.Var 3)::nil))::nil)):: (algebra.Alg.Var 1)::(algebra.Alg.Var 3)::nil))):: R_xml_0_rule_as_list_19. Definition R_xml_0_rule_as_list_21 := ((algebra.Alg.Term algebra.F.id_a__isNatIListKind ((algebra.Alg.Term algebra.F.id_nil nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_20. Definition R_xml_0_rule_as_list_22 := ((algebra.Alg.Term algebra.F.id_a__isNatIListKind ((algebra.Alg.Term algebra.F.id_zeros nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_21. Definition R_xml_0_rule_as_list_23 := ((algebra.Alg.Term algebra.F.id_a__isNatIListKind ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 1):: (algebra.Alg.Var 3)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Var 1)::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind ((algebra.Alg.Var 3)::nil))::nil)))::R_xml_0_rule_as_list_22. Definition R_xml_0_rule_as_list_24 := ((algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Term algebra.F.id_0 nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_23. Definition R_xml_0_rule_as_list_25 := ((algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Term algebra.F.id_length ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__isNatIListKind ((algebra.Alg.Var 1)::nil)))::R_xml_0_rule_as_list_24. Definition R_xml_0_rule_as_list_26 := ((algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_25. Definition R_xml_0_rule_as_list_27 := ((algebra.Alg.Term algebra.F.id_a__isNatList ((algebra.Alg.Term algebra.F.id_nil nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_26. Definition R_xml_0_rule_as_list_28 := ((algebra.Alg.Term algebra.F.id_a__isNatList ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 1):: (algebra.Alg.Var 3)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind ((algebra.Alg.Var 3)::nil))::nil)):: (algebra.Alg.Var 1)::(algebra.Alg.Var 3)::nil))):: R_xml_0_rule_as_list_27. Definition R_xml_0_rule_as_list_29 := ((algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Term algebra.F.id_nil nil)::nil)),(algebra.Alg.Term algebra.F.id_0 nil)):: R_xml_0_rule_as_list_28. Definition R_xml_0_rule_as_list_30 := ((algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 6):: (algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatList ((algebra.Alg.Var 4)::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind ((algebra.Alg.Var 4)::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isNat ((algebra.Alg.Var 6)::nil))::(algebra.Alg.Term algebra.F.id_isNatKind ((algebra.Alg.Var 6)::nil))::nil))::nil)):: (algebra.Alg.Var 4)::nil)))::R_xml_0_rule_as_list_29. Definition R_xml_0_rule_as_list_31 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_zeros nil)::nil)), (algebra.Alg.Term algebra.F.id_a__zeros nil))::R_xml_0_rule_as_list_30. Definition R_xml_0_rule_as_list_32 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::nil)))::R_xml_0_rule_as_list_31. Definition R_xml_0_rule_as_list_33 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U12 ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::nil))):: R_xml_0_rule_as_list_32. Definition R_xml_0_rule_as_list_34 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isNatList ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__isNatList ((algebra.Alg.Var 5)::nil))):: R_xml_0_rule_as_list_33. Definition R_xml_0_rule_as_list_35 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::nil)))::R_xml_0_rule_as_list_34. Definition R_xml_0_rule_as_list_36 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::nil))):: R_xml_0_rule_as_list_35. Definition R_xml_0_rule_as_list_37 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isNat ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__isNat ((algebra.Alg.Var 5)::nil))):: R_xml_0_rule_as_list_36. Definition R_xml_0_rule_as_list_38 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::nil)))::R_xml_0_rule_as_list_37. Definition R_xml_0_rule_as_list_39 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U32 ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::nil))):: R_xml_0_rule_as_list_38. Definition R_xml_0_rule_as_list_40 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8):: (algebra.Alg.Var 9)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Var 8):: (algebra.Alg.Var 9)::nil)))::R_xml_0_rule_as_list_39. Definition R_xml_0_rule_as_list_41 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::nil)))::R_xml_0_rule_as_list_40. Definition R_xml_0_rule_as_list_42 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U43 ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::nil))):: R_xml_0_rule_as_list_41. Definition R_xml_0_rule_as_list_43 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isNatIList ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__isNatIList ((algebra.Alg.Var 5)::nil))):: R_xml_0_rule_as_list_42. Definition R_xml_0_rule_as_list_44 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8):: (algebra.Alg.Var 9)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Var 8):: (algebra.Alg.Var 9)::nil)))::R_xml_0_rule_as_list_43. Definition R_xml_0_rule_as_list_45 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::nil)))::R_xml_0_rule_as_list_44. Definition R_xml_0_rule_as_list_46 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U53 ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::nil))):: R_xml_0_rule_as_list_45. Definition R_xml_0_rule_as_list_47 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::nil)))::R_xml_0_rule_as_list_46. Definition R_xml_0_rule_as_list_48 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_length ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::nil))):: R_xml_0_rule_as_list_47. Definition R_xml_0_rule_as_list_49 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Var 7)::(algebra.Alg.Var 8)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Var 8)::nil)))::R_xml_0_rule_as_list_48. Definition R_xml_0_rule_as_list_50 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isNatIListKind ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__isNatIListKind ((algebra.Alg.Var 5)::nil)))::R_xml_0_rule_as_list_49. Definition R_xml_0_rule_as_list_51 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isNatKind ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Var 5)::nil))):: R_xml_0_rule_as_list_50. Definition R_xml_0_rule_as_list_52 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 7)::(algebra.Alg.Var 8)::nil))::nil)), (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Var 8)::nil))):: R_xml_0_rule_as_list_51. Definition R_xml_0_rule_as_list_53 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_0 nil)::nil)),(algebra.Alg.Term algebra.F.id_0 nil)):: R_xml_0_rule_as_list_52. Definition R_xml_0_rule_as_list_54 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_tt nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_53. Definition R_xml_0_rule_as_list_55 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 5)::nil))::nil)))::R_xml_0_rule_as_list_54. Definition R_xml_0_rule_as_list_56 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_nil nil)::nil)),(algebra.Alg.Term algebra.F.id_nil nil)):: R_xml_0_rule_as_list_55. Definition R_xml_0_rule_as_list_57 := ((algebra.Alg.Term algebra.F.id_a__zeros nil), (algebra.Alg.Term algebra.F.id_zeros nil))::R_xml_0_rule_as_list_56. Definition R_xml_0_rule_as_list_58 := ((algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)), (algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)))::R_xml_0_rule_as_list_57. Definition R_xml_0_rule_as_list_59 := ((algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Var 5)::nil)), (algebra.Alg.Term algebra.F.id_U12 ((algebra.Alg.Var 5)::nil))):: R_xml_0_rule_as_list_58. Definition R_xml_0_rule_as_list_60 := ((algebra.Alg.Term algebra.F.id_a__isNatList ((algebra.Alg.Var 5)::nil)), (algebra.Alg.Term algebra.F.id_isNatList ((algebra.Alg.Var 5)::nil))):: R_xml_0_rule_as_list_59. Definition R_xml_0_rule_as_list_61 := ((algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)), (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)))::R_xml_0_rule_as_list_60. Definition R_xml_0_rule_as_list_62 := ((algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Var 5)::nil)), (algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Var 5)::nil))):: R_xml_0_rule_as_list_61. Definition R_xml_0_rule_as_list_63 := ((algebra.Alg.Term algebra.F.id_a__isNat ((algebra.Alg.Var 5)::nil)), (algebra.Alg.Term algebra.F.id_isNat ((algebra.Alg.Var 5)::nil))):: R_xml_0_rule_as_list_62. Definition R_xml_0_rule_as_list_64 := ((algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)), (algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)))::R_xml_0_rule_as_list_63. Definition R_xml_0_rule_as_list_65 := ((algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Var 5)::nil)), (algebra.Alg.Term algebra.F.id_U32 ((algebra.Alg.Var 5)::nil))):: R_xml_0_rule_as_list_64. Definition R_xml_0_rule_as_list_66 := ((algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)), (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))):: R_xml_0_rule_as_list_65. Definition R_xml_0_rule_as_list_67 := ((algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)), (algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)))::R_xml_0_rule_as_list_66. Definition R_xml_0_rule_as_list_68 := ((algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Var 5)::nil)), (algebra.Alg.Term algebra.F.id_U43 ((algebra.Alg.Var 5)::nil))):: R_xml_0_rule_as_list_67. Definition R_xml_0_rule_as_list_69 := ((algebra.Alg.Term algebra.F.id_a__isNatIList ((algebra.Alg.Var 5)::nil)), (algebra.Alg.Term algebra.F.id_isNatIList ((algebra.Alg.Var 5)::nil))):: R_xml_0_rule_as_list_68. Definition R_xml_0_rule_as_list_70 := ((algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)), (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))):: R_xml_0_rule_as_list_69. Definition R_xml_0_rule_as_list_71 := ((algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)), (algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)))::R_xml_0_rule_as_list_70. Definition R_xml_0_rule_as_list_72 := ((algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Var 5)::nil)), (algebra.Alg.Term algebra.F.id_U53 ((algebra.Alg.Var 5)::nil))):: R_xml_0_rule_as_list_71. Definition R_xml_0_rule_as_list_73 := ((algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)), (algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)))::R_xml_0_rule_as_list_72. Definition R_xml_0_rule_as_list_74 := ((algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Var 5)::nil)), (algebra.Alg.Term algebra.F.id_length ((algebra.Alg.Var 5)::nil))):: R_xml_0_rule_as_list_73. Definition R_xml_0_rule_as_list_75 := ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)), (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)))::R_xml_0_rule_as_list_74. Definition R_xml_0_rule_as_list_76 := ((algebra.Alg.Term algebra.F.id_a__isNatIListKind ((algebra.Alg.Var 5)::nil)), (algebra.Alg.Term algebra.F.id_isNatIListKind ((algebra.Alg.Var 5)::nil))):: R_xml_0_rule_as_list_75. Definition R_xml_0_rule_as_list_77 := ((algebra.Alg.Term algebra.F.id_a__isNatKind ((algebra.Alg.Var 5)::nil)), (algebra.Alg.Term algebra.F.id_isNatKind ((algebra.Alg.Var 5)::nil))):: R_xml_0_rule_as_list_76. Definition R_xml_0_rule_as_list := R_xml_0_rule_as_list_77. 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 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. injection H;intros ;subst;constructor 55. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. 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. injection H;intros ;subst;constructor 31. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. 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. injection H;intros ;subst;constructor 7. rewrite <- or_ext_generated.or7_equiv in H|-. case H;clear H;intros H. 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_26 (x11 x12 x13 x14 x15 x16 x17 x18 x19 x20 x21 x22 x23 x24 x25 x26 x27 x28 x29 x30 x31 x32 x33 x34 x35 x36:Prop) : Prop := | conj_26 : x11->x12->x13->x14->x15->x16->x17->x18->x19->x20->x21->x22->x23-> x24->x25->x26->x27->x28->x29->x30->x31->x32->x33->x34->x35->x36-> and_26 x11 x12 x13 x14 x15 x16 x17 x18 x19 x20 x21 x22 x23 x24 x25 x26 x27 x28 x29 x30 x31 x32 x33 x34 x35 x36 . 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_26 (forall x12, t = (algebra.Alg.Term algebra.F.id_isNatKind (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)) (forall x12 x14 x16, t = (algebra.Alg.Term algebra.F.id_U41 (x12::x14::x16::nil)) -> exists x11, exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U41 (x11::x13::x15::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x15 x16)) (forall x12 x14, t = (algebra.Alg.Term algebra.F.id_U21 (x12::x14::nil)) -> exists x11, exists x13, t' = (algebra.Alg.Term algebra.F.id_U21 (x11::x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x12 x14 x16, t = (algebra.Alg.Term algebra.F.id_U51 (x12::x14::x16::nil)) -> exists x11, exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U51 (x11::x13::x15::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x15 x16)) (forall x12, t = (algebra.Alg.Term algebra.F.id_isNatIListKind (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)) (t = (algebra.Alg.Term algebra.F.id_0 nil) -> t' = (algebra.Alg.Term algebra.F.id_0 nil)) (forall x12, t = (algebra.Alg.Term algebra.F.id_U12 (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_U12 (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)) (forall x12, t = (algebra.Alg.Term algebra.F.id_U43 (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_U43 (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)) (forall x12 x14, t = (algebra.Alg.Term algebra.F.id_U31 (x12::x14::nil)) -> exists x11, exists x13, t' = (algebra.Alg.Term algebra.F.id_U31 (x11::x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x12, t = (algebra.Alg.Term algebra.F.id_U53 (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_U53 (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)) (forall x12 x14, t = (algebra.Alg.Term algebra.F.id_and (x12::x14::nil)) -> exists x11, exists x13, t' = (algebra.Alg.Term algebra.F.id_and (x11::x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x12 x14, t = (algebra.Alg.Term algebra.F.id_cons (x12::x14::nil)) -> exists x11, exists x13, t' = (algebra.Alg.Term algebra.F.id_cons (x11::x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x12 x14, t = (algebra.Alg.Term algebra.F.id_U11 (x12::x14::nil)) -> exists x11, exists x13, t' = (algebra.Alg.Term algebra.F.id_U11 (x11::x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x12 x14, t = (algebra.Alg.Term algebra.F.id_U42 (x12::x14::nil)) -> exists x11, exists x13, t' = (algebra.Alg.Term algebra.F.id_U42 (x11::x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x12, t = (algebra.Alg.Term algebra.F.id_length (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_length (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)) (t = (algebra.Alg.Term algebra.F.id_tt nil) -> t' = (algebra.Alg.Term algebra.F.id_tt nil)) (forall x12, t = (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_U22 (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)) (forall x12, t = (algebra.Alg.Term algebra.F.id_s (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_s (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)) (forall x12 x14, t = (algebra.Alg.Term algebra.F.id_U52 (x12::x14::nil)) -> exists x11, exists x13, t' = (algebra.Alg.Term algebra.F.id_U52 (x11::x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (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_zeros nil) -> t' = (algebra.Alg.Term algebra.F.id_zeros nil)) (forall x12, t = (algebra.Alg.Term algebra.F.id_isNatList (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_isNatList (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)) (forall x12, t = (algebra.Alg.Term algebra.F.id_isNatIList (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)) (forall x12, t = (algebra.Alg.Term algebra.F.id_U32 (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_U32 (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)) (forall x12 x14, t = (algebra.Alg.Term algebra.F.id_U61 (x12::x14::nil)) -> exists x11, exists x13, t' = (algebra.Alg.Term algebra.F.id_U61 (x11::x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x12, t = (algebra.Alg.Term algebra.F.id_isNat (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_isNat (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)). Proof. intros t t' H. induction H as [|y IH z z_to_y] using closure_extension.refl_trans_clos_ind2. constructor 1. intros x12 H;exists x12;intuition;constructor 1. intros x12 x14 x16 H;exists x12;exists x14;exists x16;intuition; constructor 1. intros x12 x14 H;exists x12;exists x14;intuition;constructor 1. intros x12 x14 x16 H;exists x12;exists x14;exists x16;intuition; constructor 1. intros x12 H;exists x12;intuition;constructor 1. intros H;intuition;constructor 1. intros x12 H;exists x12;intuition;constructor 1. intros x12 H;exists x12;intuition;constructor 1. intros x12 x14 H;exists x12;exists x14;intuition;constructor 1. intros x12 H;exists x12;intuition;constructor 1. intros x12 x14 H;exists x12;exists x14;intuition;constructor 1. intros x12 x14 H;exists x12;exists x14;intuition;constructor 1. intros x12 x14 H;exists x12;exists x14;intuition;constructor 1. intros x12 x14 H;exists x12;exists x14;intuition;constructor 1. intros x12 H;exists x12;intuition;constructor 1. intros H;intuition;constructor 1. intros x12 H;exists x12;intuition;constructor 1. intros x12 H;exists x12;intuition;constructor 1. intros x12 x14 H;exists x12;exists x14;intuition;constructor 1. intros H;intuition;constructor 1. intros H;intuition;constructor 1. intros x12 H;exists x12;intuition;constructor 1. intros x12 H;exists x12;intuition;constructor 1. intros x12 H;exists x12;intuition;constructor 1. intros x12 x14 H;exists x12;exists x14;intuition;constructor 1. intros x12 H;exists x12;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_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat]. constructor. clear H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat;intros x12 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 x12 |- _ => destruct (H_id_isNatKind y (refl_equal _)) as [x11];intros ;intuition; exists x11;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_isNatKind H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat; intros x12 x14 x16 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 x12 |- _ => destruct (H_id_U41 y x14 x16 (refl_equal _)) as [x11 [x13 [x15]]]; intros ;intuition;exists x11;exists x13;exists x15;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x14 |- _ => destruct (H_id_U41 x12 y x16 (refl_equal _)) as [x11 [x13 [x15]]]; intros ;intuition;exists x11;exists x13;exists x15;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x16 |- _ => destruct (H_id_U41 x12 x14 y (refl_equal _)) as [x11 [x13 [x15]]]; intros ;intuition;exists x11;exists x13;exists x15;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_isNatKind H_id_U41 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat; intros x12 x14 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 x12 |- _ => destruct (H_id_U21 y x14 (refl_equal _)) as [x11 [x13]];intros ; intuition;exists x11;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x14 |- _ => destruct (H_id_U21 x12 y (refl_equal _)) as [x11 [x13]];intros ; intuition;exists x11;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat; intros x12 x14 x16 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 x12 |- _ => destruct (H_id_U51 y x14 x16 (refl_equal _)) as [x11 [x13 [x15]]]; intros ;intuition;exists x11;exists x13;exists x15;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x14 |- _ => destruct (H_id_U51 x12 y x16 (refl_equal _)) as [x11 [x13 [x15]]]; intros ;intuition;exists x11;exists x13;exists x15;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x16 |- _ => destruct (H_id_U51 x12 x14 y (refl_equal _)) as [x11 [x13 [x15]]]; intros ;intuition;exists x11;exists x13;exists x15;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat;intros x12 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 x12 |- _ => destruct (H_id_isNatIListKind y (refl_equal _)) as [x11];intros ; intuition;exists x11;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat;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_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat;intros x12 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 x12 |- _ => destruct (H_id_U12 y (refl_equal _)) as [x11];intros ;intuition; exists x11;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat;intros x12 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 x12 |- _ => destruct (H_id_U43 y (refl_equal _)) as [x11];intros ;intuition; exists x11;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat; intros x12 x14 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 x12 |- _ => destruct (H_id_U31 y x14 (refl_equal _)) as [x11 [x13]];intros ; intuition;exists x11;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x14 |- _ => destruct (H_id_U31 x12 y (refl_equal _)) as [x11 [x13]];intros ; intuition;exists x11;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat;intros x12 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 x12 |- _ => destruct (H_id_U53 y (refl_equal _)) as [x11];intros ;intuition; exists x11;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat; intros x12 x14 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 x12 |- _ => destruct (H_id_and y x14 (refl_equal _)) as [x11 [x13]];intros ; intuition;exists x11;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x14 |- _ => destruct (H_id_and x12 y (refl_equal _)) as [x11 [x13]];intros ; intuition;exists x11;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat; intros x12 x14 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 x12 |- _ => destruct (H_id_cons y x14 (refl_equal _)) as [x11 [x13]];intros ; intuition;exists x11;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x14 |- _ => destruct (H_id_cons x12 y (refl_equal _)) as [x11 [x13]];intros ; intuition;exists x11;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat; intros x12 x14 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 x12 |- _ => destruct (H_id_U11 y x14 (refl_equal _)) as [x11 [x13]];intros ; intuition;exists x11;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x14 |- _ => destruct (H_id_U11 x12 y (refl_equal _)) as [x11 [x13]];intros ; intuition;exists x11;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat; intros x12 x14 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 x12 |- _ => destruct (H_id_U42 y x14 (refl_equal _)) as [x11 [x13]];intros ; intuition;exists x11;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x14 |- _ => destruct (H_id_U42 x12 y (refl_equal _)) as [x11 [x13]];intros ; intuition;exists x11;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat;intros x12 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 x12 |- _ => destruct (H_id_length y (refl_equal _)) as [x11];intros ;intuition; exists x11;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat;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_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat;intros x12 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 x12 |- _ => destruct (H_id_U22 y (refl_equal _)) as [x11];intros ;intuition; exists x11;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat;intros x12 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 x12 |- _ => destruct (H_id_s y (refl_equal _)) as [x11];intros ;intuition; exists x11;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat; intros x12 x14 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 x12 |- _ => destruct (H_id_U52 y x14 (refl_equal _)) as [x11 [x13]];intros ; intuition;exists x11;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x14 |- _ => destruct (H_id_U52 x12 y (refl_equal _)) as [x11 [x13]];intros ; intuition;exists x11;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat;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_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat;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_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatIList H_id_U32 H_id_U61 H_id_isNat;intros x12 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 x12 |- _ => destruct (H_id_isNatList y (refl_equal _)) as [x11];intros ;intuition; exists x11;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_U32 H_id_U61 H_id_isNat;intros x12 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 x12 |- _ => destruct (H_id_isNatIList y (refl_equal _)) as [x11];intros ;intuition; exists x11;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U61 H_id_isNat;intros x12 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 x12 |- _ => destruct (H_id_U32 y (refl_equal _)) as [x11];intros ;intuition; exists x11;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_isNat;intros x12 x14 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 x12 |- _ => destruct (H_id_U61 y x14 (refl_equal _)) as [x11 [x13]];intros ; intuition;exists x11;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x14 |- _ => destruct (H_id_U61 x12 y (refl_equal _)) as [x11 [x13]];intros ; intuition;exists x11;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_isNatKind H_id_U41 H_id_U21 H_id_U51 H_id_isNatIListKind H_id_0 H_id_U12 H_id_U43 H_id_U31 H_id_U53 H_id_and H_id_cons H_id_U11 H_id_U42 H_id_length H_id_tt H_id_U22 H_id_s H_id_U52 H_id_nil H_id_zeros H_id_isNatList H_id_isNatIList H_id_U32 H_id_U61;intros x12 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 x12 |- _ => destruct (H_id_isNat y (refl_equal _)) as [x11];intros ;intuition; exists x11;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . Qed. Lemma id_isNatKind_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12, t = (algebra.Alg.Term algebra.F.id_isNatKind (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U41_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12 x14 x16, t = (algebra.Alg.Term algebra.F.id_U41 (x12::x14::x16::nil)) -> exists x11, exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U41 (x11::x13::x15::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x15 x16). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U21_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12 x14, t = (algebra.Alg.Term algebra.F.id_U21 (x12::x14::nil)) -> exists x11, exists x13, t' = (algebra.Alg.Term algebra.F.id_U21 (x11::x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U51_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12 x14 x16, t = (algebra.Alg.Term algebra.F.id_U51 (x12::x14::x16::nil)) -> exists x11, exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U51 (x11::x13::x15::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x15 x16). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_isNatIListKind_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12, t = (algebra.Alg.Term algebra.F.id_isNatIListKind (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_0_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> t = (algebra.Alg.Term algebra.F.id_0 nil) -> t' = (algebra.Alg.Term algebra.F.id_0 nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U12_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12, t = (algebra.Alg.Term algebra.F.id_U12 (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_U12 (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U43_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12, t = (algebra.Alg.Term algebra.F.id_U43 (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_U43 (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U31_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12 x14, t = (algebra.Alg.Term algebra.F.id_U31 (x12::x14::nil)) -> exists x11, exists x13, t' = (algebra.Alg.Term algebra.F.id_U31 (x11::x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U53_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12, t = (algebra.Alg.Term algebra.F.id_U53 (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_U53 (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_and_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12 x14, t = (algebra.Alg.Term algebra.F.id_and (x12::x14::nil)) -> exists x11, exists x13, t' = (algebra.Alg.Term algebra.F.id_and (x11::x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_cons_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12 x14, t = (algebra.Alg.Term algebra.F.id_cons (x12::x14::nil)) -> exists x11, exists x13, t' = (algebra.Alg.Term algebra.F.id_cons (x11::x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U11_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12 x14, t = (algebra.Alg.Term algebra.F.id_U11 (x12::x14::nil)) -> exists x11, exists x13, t' = (algebra.Alg.Term algebra.F.id_U11 (x11::x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U42_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12 x14, t = (algebra.Alg.Term algebra.F.id_U42 (x12::x14::nil)) -> exists x11, exists x13, t' = (algebra.Alg.Term algebra.F.id_U42 (x11::x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_length_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12, t = (algebra.Alg.Term algebra.F.id_length (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_length (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12). 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_U22_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12, t = (algebra.Alg.Term algebra.F.id_U22 (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_U22 (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_s_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12, t = (algebra.Alg.Term algebra.F.id_s (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_s (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U52_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12 x14, t = (algebra.Alg.Term algebra.F.id_U52 (x12::x14::nil)) -> exists x11, exists x13, t' = (algebra.Alg.Term algebra.F.id_U52 (x11::x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14). 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_zeros_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_zeros nil) -> t' = (algebra.Alg.Term algebra.F.id_zeros nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_isNatList_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12, t = (algebra.Alg.Term algebra.F.id_isNatList (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_isNatList (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_isNatIList_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12, t = (algebra.Alg.Term algebra.F.id_isNatIList (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U32_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12, t = (algebra.Alg.Term algebra.F.id_U32 (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_U32 (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U61_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12 x14, t = (algebra.Alg.Term algebra.F.id_U61 (x12::x14::nil)) -> exists x11, exists x13, t' = (algebra.Alg.Term algebra.F.id_U61 (x11::x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_isNat_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x12, t = (algebra.Alg.Term algebra.F.id_isNat (x12::nil)) -> exists x11, t' = (algebra.Alg.Term algebra.F.id_isNat (x11::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x11 x12). 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_isNatKind (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isNatKind_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_U41 (?x13::?x12::?x11::nil)) |- _ => let x13 := fresh "x" in (let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (let Hred3 := fresh "Hred" in (destruct (id_U41_is_R_xml_0_constructor H (refl_equal _)) as [x13 [x12 [x11 [Heq [Hred3 [Hred2 Hred1]]]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U21 (?x12::?x11::nil)) |- _ => let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_U21_is_R_xml_0_constructor H (refl_equal _)) as [x12 [x11 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U51 (?x13::?x12::?x11::nil)) |- _ => let x13 := fresh "x" in (let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (let Hred3 := fresh "Hred" in (destruct (id_U51_is_R_xml_0_constructor H (refl_equal _)) as [x13 [x12 [x11 [Heq [Hred3 [Hred2 Hred1]]]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_isNatIListKind (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isNatIListKind_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_0 nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_0_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; impossible_star_reduction_R_xml_0 )) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U12 (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_U12_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_U43 (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_U43_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_U31 (?x12::?x11::nil)) |- _ => let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_U31_is_R_xml_0_constructor H (refl_equal _)) as [x12 [x11 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U53 (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_U53_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_and (?x12::?x11::nil)) |- _ => let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_and_is_R_xml_0_constructor H (refl_equal _)) as [x12 [x11 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_cons (?x12::?x11::nil)) |- _ => let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_cons_is_R_xml_0_constructor H (refl_equal _)) as [x12 [x11 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U11 (?x12::?x11::nil)) |- _ => let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_U11_is_R_xml_0_constructor H (refl_equal _)) as [x12 [x11 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U42 (?x12::?x11::nil)) |- _ => let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_U42_is_R_xml_0_constructor H (refl_equal _)) as [x12 [x11 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_length (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_length_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_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_U22 (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_U22_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_s (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_s_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_U52 (?x12::?x11::nil)) |- _ => let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_U52_is_R_xml_0_constructor H (refl_equal _)) as [x12 [x11 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_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_zeros nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_zeros_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_isNatList (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isNatList_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_isNatIList (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isNatIList_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_U32 (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_U32_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_U61 (?x12::?x11::nil)) |- _ => let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_U61_is_R_xml_0_constructor H (refl_equal _)) as [x12 [x11 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_isNat (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isNat_is_R_xml_0_constructor H (refl_equal _)) as [x11 [Heq Hred1]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))) end . Ltac simplify_star_reduction_R_xml_0 := match goal with | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_isNatKind (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isNatKind_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_U41 (?x13::?x12::?x11::nil)) |- _ => let x13 := fresh "x" in (let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (let Hred3 := fresh "Hred" in (destruct (id_U41_is_R_xml_0_constructor H (refl_equal _)) as [x13 [x12 [x11 [Heq [Hred3 [Hred2 Hred1]]]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U21 (?x12::?x11::nil)) |- _ => let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_U21_is_R_xml_0_constructor H (refl_equal _)) as [x12 [x11 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U51 (?x13::?x12::?x11::nil)) |- _ => let x13 := fresh "x" in (let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (let Hred3 := fresh "Hred" in (destruct (id_U51_is_R_xml_0_constructor H (refl_equal _)) as [x13 [x12 [x11 [Heq [Hred3 [Hred2 Hred1]]]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_isNatIListKind (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isNatIListKind_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_0 nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_0_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; try (simplify_star_reduction_R_xml_0 ))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U12 (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_U12_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_U43 (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_U43_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_U31 (?x12::?x11::nil)) |- _ => let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_U31_is_R_xml_0_constructor H (refl_equal _)) as [x12 [x11 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U53 (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_U53_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_and (?x12::?x11::nil)) |- _ => let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_and_is_R_xml_0_constructor H (refl_equal _)) as [x12 [x11 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_cons (?x12::?x11::nil)) |- _ => let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_cons_is_R_xml_0_constructor H (refl_equal _)) as [x12 [x11 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U11 (?x12::?x11::nil)) |- _ => let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_U11_is_R_xml_0_constructor H (refl_equal _)) as [x12 [x11 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U42 (?x12::?x11::nil)) |- _ => let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_U42_is_R_xml_0_constructor H (refl_equal _)) as [x12 [x11 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_length (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_length_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_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_U22 (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_U22_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_s (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_s_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_U52 (?x12::?x11::nil)) |- _ => let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_U52_is_R_xml_0_constructor H (refl_equal _)) as [x12 [x11 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_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_zeros nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_zeros_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_isNatList (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isNatList_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_isNatIList (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isNatIList_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_U32 (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_U32_is_R_xml_0_constructor H (refl_equal _)) as [x11 [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_U61 (?x12::?x11::nil)) |- _ => let x12 := fresh "x" in (let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_U61_is_R_xml_0_constructor H (refl_equal _)) as [x12 [x11 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_isNat (?x11::nil)) |- _ => let x11 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isNat_is_R_xml_0_constructor H (refl_equal _)) as [x11 [Heq Hred1]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))) end . End R_xml_0_deep_rew. Module InterpGen := interp.Interp(algebra.EQT). Module ddp := dp.MakeDP(algebra.EQT). Module SymbType. Definition A := algebra.Alg.F.Symb.A. End SymbType. Module Symb_more_list := more_list_extention.Make(SymbType)(algebra.Alg.F.Symb). Module SymbSet := list_set.Make(algebra.F.Symb). Module Interp. Section S. Require Import interp. Hypothesis A : Type. Hypothesis Ale Alt Aeq : A -> A -> Prop. Hypothesis Aop : interp.ordering_pair Aeq Alt Ale. Hypothesis A0 : A. Notation Local "a <= b" := (Ale a b). Hypothesis P_id_a__zeros : A. Hypothesis P_id_isNatKind : A ->A. Hypothesis P_id_a__isNatIList : A ->A. Hypothesis P_id_a__U21 : A ->A ->A. Hypothesis P_id_U41 : A ->A ->A ->A. Hypothesis P_id_a__and : A ->A ->A. Hypothesis P_id_a__U11 : A ->A ->A. Hypothesis P_id_U21 : A ->A ->A. Hypothesis P_id_a__U61 : A ->A ->A. Hypothesis P_id_a__U32 : A ->A. Hypothesis P_id_U51 : A ->A ->A ->A. Hypothesis P_id_isNatIListKind : A ->A. Hypothesis P_id_0 : A. Hypothesis P_id_U12 : A ->A. Hypothesis P_id_a__U52 : A ->A ->A. Hypothesis P_id_a__isNat : A ->A. Hypothesis P_id_U43 : A ->A. Hypothesis P_id_a__isNatIListKind : A ->A. Hypothesis P_id_a__U12 : A ->A. Hypothesis P_id_U31 : A ->A ->A. Hypothesis P_id_a__length : A ->A. Hypothesis P_id_a__U42 : A ->A ->A. Hypothesis P_id_U53 : A ->A. Hypothesis P_id_and : A ->A ->A. Hypothesis P_id_cons : A ->A ->A. Hypothesis P_id_U11 : A ->A ->A. Hypothesis P_id_a__U51 : A ->A ->A ->A. Hypothesis P_id_a__U22 : A ->A. Hypothesis P_id_U42 : A ->A ->A. Hypothesis P_id_length : A ->A. Hypothesis P_id_tt : A. Hypothesis P_id_U22 : A ->A. Hypothesis P_id_s : A ->A. Hypothesis P_id_a__U41 : A ->A ->A ->A. Hypothesis P_id_U52 : A ->A ->A. Hypothesis P_id_nil : A. Hypothesis P_id_zeros : A. Hypothesis P_id_isNatList : A ->A. Hypothesis P_id_a__U53 : A ->A. Hypothesis P_id_a__U31 : A ->A ->A. Hypothesis P_id_isNatIList : A ->A. Hypothesis P_id_a__isNatKind : A ->A. Hypothesis P_id_a__isNatList : A ->A. Hypothesis P_id_U32 : A ->A. Hypothesis P_id_mark : A ->A. Hypothesis P_id_a__U43 : A ->A. Hypothesis P_id_U61 : A ->A ->A. Hypothesis P_id_isNat : A ->A. Hypothesis P_id_isNatKind_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_isNatKind x12 <= P_id_isNatKind x11. Hypothesis P_id_a__isNatIList_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Hypothesis P_id_a__U21_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Hypothesis P_id_U41_monotonic : forall x16 x12 x14 x13 x11 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Hypothesis P_id_a__and_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_a__and x12 x14 <= P_id_a__and x11 x13. Hypothesis P_id_a__U11_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Hypothesis P_id_U21_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_U21 x12 x14 <= P_id_U21 x11 x13. Hypothesis P_id_a__U61_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Hypothesis P_id_a__U32_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_a__U32 x12 <= P_id_a__U32 x11. Hypothesis P_id_U51_monotonic : forall x16 x12 x14 x13 x11 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Hypothesis P_id_isNatIListKind_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Hypothesis P_id_U12_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Hypothesis P_id_a__U52_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Hypothesis P_id_a__isNat_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_a__isNat x12 <= P_id_a__isNat x11. Hypothesis P_id_U43_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Hypothesis P_id_a__isNatIListKind_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Hypothesis P_id_a__U12_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_a__U12 x12 <= P_id_a__U12 x11. Hypothesis P_id_U31_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_U31 x12 x14 <= P_id_U31 x11 x13. Hypothesis P_id_a__length_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_a__length x12 <= P_id_a__length x11. Hypothesis P_id_a__U42_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Hypothesis P_id_U53_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Hypothesis P_id_and_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_and x12 x14 <= P_id_and x11 x13. Hypothesis P_id_cons_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_cons x12 x14 <= P_id_cons x11 x13. Hypothesis P_id_U11_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_U11 x12 x14 <= P_id_U11 x11 x13. Hypothesis P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Hypothesis P_id_a__U22_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_a__U22 x12 <= P_id_a__U22 x11. Hypothesis P_id_U42_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_U42 x12 x14 <= P_id_U42 x11 x13. Hypothesis P_id_length_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_length x12 <= P_id_length x11. Hypothesis P_id_U22_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Hypothesis P_id_s_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Hypothesis P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Hypothesis P_id_U52_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_U52 x12 x14 <= P_id_U52 x11 x13. Hypothesis P_id_isNatList_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_isNatList x12 <= P_id_isNatList x11. Hypothesis P_id_a__U53_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_a__U53 x12 <= P_id_a__U53 x11. Hypothesis P_id_a__U31_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Hypothesis P_id_isNatIList_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_isNatIList x12 <= P_id_isNatIList x11. Hypothesis P_id_a__isNatKind_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Hypothesis P_id_a__isNatList_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Hypothesis P_id_U32_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Hypothesis P_id_mark_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Hypothesis P_id_a__U43_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_a__U43 x12 <= P_id_a__U43 x11. Hypothesis P_id_U61_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_U61 x12 x14 <= P_id_U61 x11 x13. Hypothesis P_id_isNat_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_isNat x12 <= P_id_isNat x11. Hypothesis P_id_a__zeros_bounded : A0 <= P_id_a__zeros . Hypothesis P_id_isNatKind_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_isNatKind x11. Hypothesis P_id_a__isNatIList_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_a__isNatIList x11. Hypothesis P_id_a__U21_bounded : forall x12 x11, (A0 <= x11) ->(A0 <= x12) ->A0 <= P_id_a__U21 x12 x11. Hypothesis P_id_U41_bounded : forall x12 x13 x11, (A0 <= x11) ->(A0 <= x12) ->(A0 <= x13) ->A0 <= P_id_U41 x13 x12 x11. Hypothesis P_id_a__and_bounded : forall x12 x11, (A0 <= x11) ->(A0 <= x12) ->A0 <= P_id_a__and x12 x11. Hypothesis P_id_a__U11_bounded : forall x12 x11, (A0 <= x11) ->(A0 <= x12) ->A0 <= P_id_a__U11 x12 x11. Hypothesis P_id_U21_bounded : forall x12 x11, (A0 <= x11) ->(A0 <= x12) ->A0 <= P_id_U21 x12 x11. Hypothesis P_id_a__U61_bounded : forall x12 x11, (A0 <= x11) ->(A0 <= x12) ->A0 <= P_id_a__U61 x12 x11. Hypothesis P_id_a__U32_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_a__U32 x11. Hypothesis P_id_U51_bounded : forall x12 x13 x11, (A0 <= x11) ->(A0 <= x12) ->(A0 <= x13) ->A0 <= P_id_U51 x13 x12 x11. Hypothesis P_id_isNatIListKind_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_isNatIListKind x11. Hypothesis P_id_0_bounded : A0 <= P_id_0 . Hypothesis P_id_U12_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_U12 x11. Hypothesis P_id_a__U52_bounded : forall x12 x11, (A0 <= x11) ->(A0 <= x12) ->A0 <= P_id_a__U52 x12 x11. Hypothesis P_id_a__isNat_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_a__isNat x11. Hypothesis P_id_U43_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_U43 x11. Hypothesis P_id_a__isNatIListKind_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_a__isNatIListKind x11. Hypothesis P_id_a__U12_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_a__U12 x11. Hypothesis P_id_U31_bounded : forall x12 x11, (A0 <= x11) ->(A0 <= x12) ->A0 <= P_id_U31 x12 x11. Hypothesis P_id_a__length_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_a__length x11. Hypothesis P_id_a__U42_bounded : forall x12 x11, (A0 <= x11) ->(A0 <= x12) ->A0 <= P_id_a__U42 x12 x11. Hypothesis P_id_U53_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_U53 x11. Hypothesis P_id_and_bounded : forall x12 x11, (A0 <= x11) ->(A0 <= x12) ->A0 <= P_id_and x12 x11. Hypothesis P_id_cons_bounded : forall x12 x11, (A0 <= x11) ->(A0 <= x12) ->A0 <= P_id_cons x12 x11. Hypothesis P_id_U11_bounded : forall x12 x11, (A0 <= x11) ->(A0 <= x12) ->A0 <= P_id_U11 x12 x11. Hypothesis P_id_a__U51_bounded : forall x12 x13 x11, (A0 <= x11) ->(A0 <= x12) ->(A0 <= x13) ->A0 <= P_id_a__U51 x13 x12 x11. Hypothesis P_id_a__U22_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_a__U22 x11. Hypothesis P_id_U42_bounded : forall x12 x11, (A0 <= x11) ->(A0 <= x12) ->A0 <= P_id_U42 x12 x11. Hypothesis P_id_length_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_length x11. Hypothesis P_id_tt_bounded : A0 <= P_id_tt . Hypothesis P_id_U22_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_U22 x11. Hypothesis P_id_s_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_s x11. Hypothesis P_id_a__U41_bounded : forall x12 x13 x11, (A0 <= x11) ->(A0 <= x12) ->(A0 <= x13) ->A0 <= P_id_a__U41 x13 x12 x11. Hypothesis P_id_U52_bounded : forall x12 x11, (A0 <= x11) ->(A0 <= x12) ->A0 <= P_id_U52 x12 x11. Hypothesis P_id_nil_bounded : A0 <= P_id_nil . Hypothesis P_id_zeros_bounded : A0 <= P_id_zeros . Hypothesis P_id_isNatList_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_isNatList x11. Hypothesis P_id_a__U53_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_a__U53 x11. Hypothesis P_id_a__U31_bounded : forall x12 x11, (A0 <= x11) ->(A0 <= x12) ->A0 <= P_id_a__U31 x12 x11. Hypothesis P_id_isNatIList_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_isNatIList x11. Hypothesis P_id_a__isNatKind_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_a__isNatKind x11. Hypothesis P_id_a__isNatList_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_a__isNatList x11. Hypothesis P_id_U32_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_U32 x11. Hypothesis P_id_mark_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_mark x11. Hypothesis P_id_a__U43_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_a__U43 x11. Hypothesis P_id_U61_bounded : forall x12 x11, (A0 <= x11) ->(A0 <= x12) ->A0 <= P_id_U61 x12 x11. Hypothesis P_id_isNat_bounded : forall x11, (A0 <= x11) ->A0 <= P_id_isNat x11. Fixpoint measure t { struct t } := match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12::x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12::x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12::x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12::x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12::x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12::x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12::x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12::x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => A0 end. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12::x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12::x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12::x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12::x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12::x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12::x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12::x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12::x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_a__zeros => P_id_a__zeros | algebra.F.id_isNatKind => P_id_isNatKind | algebra.F.id_a__isNatIList => P_id_a__isNatIList | algebra.F.id_a__U21 => P_id_a__U21 | algebra.F.id_U41 => P_id_U41 | algebra.F.id_a__and => P_id_a__and | algebra.F.id_a__U11 => P_id_a__U11 | algebra.F.id_U21 => P_id_U21 | algebra.F.id_a__U61 => P_id_a__U61 | algebra.F.id_a__U32 => P_id_a__U32 | algebra.F.id_U51 => P_id_U51 | algebra.F.id_isNatIListKind => P_id_isNatIListKind | algebra.F.id_0 => P_id_0 | algebra.F.id_U12 => P_id_U12 | algebra.F.id_a__U52 => P_id_a__U52 | algebra.F.id_a__isNat => P_id_a__isNat | algebra.F.id_U43 => P_id_U43 | algebra.F.id_a__isNatIListKind => P_id_a__isNatIListKind | algebra.F.id_a__U12 => P_id_a__U12 | algebra.F.id_U31 => P_id_U31 | algebra.F.id_a__length => P_id_a__length | algebra.F.id_a__U42 => P_id_a__U42 | algebra.F.id_U53 => P_id_U53 | algebra.F.id_and => P_id_and | algebra.F.id_cons => P_id_cons | algebra.F.id_U11 => P_id_U11 | algebra.F.id_a__U51 => P_id_a__U51 | algebra.F.id_a__U22 => P_id_a__U22 | algebra.F.id_U42 => P_id_U42 | algebra.F.id_length => P_id_length | algebra.F.id_tt => P_id_tt | algebra.F.id_U22 => P_id_U22 | algebra.F.id_s => P_id_s | algebra.F.id_a__U41 => P_id_a__U41 | algebra.F.id_U52 => P_id_U52 | algebra.F.id_nil => P_id_nil | algebra.F.id_zeros => P_id_zeros | algebra.F.id_isNatList => P_id_isNatList | algebra.F.id_a__U53 => P_id_a__U53 | algebra.F.id_a__U31 => P_id_a__U31 | algebra.F.id_isNatIList => P_id_isNatIList | algebra.F.id_a__isNatKind => P_id_a__isNatKind | algebra.F.id_a__isNatList => P_id_a__isNatList | algebra.F.id_U32 => P_id_U32 | algebra.F.id_mark => P_id_mark | algebra.F.id_a__U43 => P_id_a__U43 | algebra.F.id_U61 => P_id_U61 | algebra.F.id_isNat => P_id_isNat 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_a__zeros => match l with | nil => _ | _::_ => _ end | algebra.F.id_isNatKind => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__isNatIList => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U21 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U41 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_a__and => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U11 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U21 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U61 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U32 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U51 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_isNatIListKind => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_0 => match l with | nil => _ | _::_ => _ end | algebra.F.id_U12 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U52 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__isNat => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U43 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__isNatIListKind => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U12 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U31 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__length => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U42 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U53 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_and => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_cons => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U11 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U51 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_a__U22 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U42 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_length => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_tt => match l with | nil => _ | _::_ => _ end | algebra.F.id_U22 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_s => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U41 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_U52 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_nil => match l with | nil => _ | _::_ => _ end | algebra.F.id_zeros => match l with | nil => _ | _::_ => _ end | algebra.F.id_isNatList => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U53 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U31 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_isNatIList => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__isNatKind => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__isNatList => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U32 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_mark => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U43 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U61 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_isNat => 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_a__zeros_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNatKind_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatIList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U21_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U41_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__and_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U11_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U21_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U61_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U32_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U51_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNatIListKind_bounded;assumption. vm_compute in |-*;intros ;apply P_id_0_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U12_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U52_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNat_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U43_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatIListKind_bounded; assumption. vm_compute in |-*;intros ;apply P_id_a__U12_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U31_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__length_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U42_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U53_bounded;assumption. vm_compute in |-*;intros ;apply P_id_and_bounded;assumption. vm_compute in |-*;intros ;apply P_id_cons_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U11_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U51_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U22_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U42_bounded;assumption. vm_compute in |-*;intros ;apply P_id_length_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tt_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U22_bounded;assumption. vm_compute in |-*;intros ;apply P_id_s_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U41_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U52_bounded;assumption. vm_compute in |-*;intros ;apply P_id_nil_bounded;assumption. vm_compute in |-*;intros ;apply P_id_zeros_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNatList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U53_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U31_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNatIList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatKind_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U32_bounded;assumption. vm_compute in |-*;intros ;apply P_id_mark_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U43_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U61_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNat_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 |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_isNatKind_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatIList_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U21_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U41_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__and_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U11_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U21_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U61_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U32_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U51_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isNatIListKind_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_U12_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U52_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isNat_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U43_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatIListKind_monotonic; assumption. vm_compute in |-*;intros ;apply P_id_a__U12_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U31_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__length_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U42_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U53_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_and_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_cons_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U11_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U51_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U22_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U42_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_length_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_U22_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_s_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U41_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U52_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_isNatList_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U53_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U31_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isNatIList_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatKind_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatList_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U32_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_mark_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U43_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U61_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isNat_monotonic;assumption. intros f. case f. vm_compute in |-*;intros ;apply P_id_a__zeros_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNatKind_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatIList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U21_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U41_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__and_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U11_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U21_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U61_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U32_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U51_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNatIListKind_bounded;assumption. vm_compute in |-*;intros ;apply P_id_0_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U12_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U52_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNat_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U43_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatIListKind_bounded; assumption. vm_compute in |-*;intros ;apply P_id_a__U12_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U31_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__length_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U42_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U53_bounded;assumption. vm_compute in |-*;intros ;apply P_id_and_bounded;assumption. vm_compute in |-*;intros ;apply P_id_cons_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U11_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U51_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U22_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U42_bounded;assumption. vm_compute in |-*;intros ;apply P_id_length_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tt_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U22_bounded;assumption. vm_compute in |-*;intros ;apply P_id_s_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U41_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U52_bounded;assumption. vm_compute in |-*;intros ;apply P_id_nil_bounded;assumption. vm_compute in |-*;intros ;apply P_id_zeros_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNatList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U53_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U31_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNatIList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatKind_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U32_bounded;assumption. vm_compute in |-*;intros ;apply P_id_mark_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U43_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U61_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNat_bounded;assumption. intros . do 2 (rewrite <- same_measure in |-*). apply rules_monotonic;assumption. assumption. Qed. Hypothesis P_id_A__LENGTH : A ->A. Hypothesis P_id_A__U11 : A ->A ->A. Hypothesis P_id_A__U41 : A ->A ->A ->A. Hypothesis P_id_A__ISNATKIND : A ->A. Hypothesis P_id_A__U22 : A ->A. Hypothesis P_id_A__U51 : A ->A ->A ->A. Hypothesis P_id_A__AND : A ->A ->A. Hypothesis P_id_A__ISNATLIST : A ->A. Hypothesis P_id_A__U43 : A ->A. Hypothesis P_id_A__U31 : A ->A ->A. Hypothesis P_id_A__U53 : A ->A. Hypothesis P_id_MARK : A ->A. Hypothesis P_id_A__U12 : A ->A. Hypothesis P_id_A__U42 : A ->A ->A. Hypothesis P_id_A__ZEROS : A. Hypothesis P_id_A__ISNAT : A ->A. Hypothesis P_id_A__U52 : A ->A ->A. Hypothesis P_id_A__ISNATILISTKIND : A ->A. Hypothesis P_id_A__U21 : A ->A ->A. Hypothesis P_id_A__ISNATILIST : A ->A. Hypothesis P_id_A__U32 : A ->A. Hypothesis P_id_A__U61 : A ->A ->A. Hypothesis P_id_A__LENGTH_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Hypothesis P_id_A__U11_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Hypothesis P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Hypothesis P_id_A__ISNATKIND_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Hypothesis P_id_A__U22_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_A__U22 x12 <= P_id_A__U22 x11. Hypothesis P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Hypothesis P_id_A__AND_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Hypothesis P_id_A__ISNATLIST_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Hypothesis P_id_A__U43_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_A__U43 x12 <= P_id_A__U43 x11. Hypothesis P_id_A__U31_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Hypothesis P_id_A__U53_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_A__U53 x12 <= P_id_A__U53 x11. Hypothesis P_id_MARK_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Hypothesis P_id_A__U12_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_A__U12 x12 <= P_id_A__U12 x11. Hypothesis P_id_A__U42_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Hypothesis P_id_A__ISNAT_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Hypothesis P_id_A__U52_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Hypothesis P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Hypothesis P_id_A__U21_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Hypothesis P_id_A__ISNATILIST_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Hypothesis P_id_A__U32_monotonic : forall x12 x11, (A0 <= x12)/\ (x12 <= x11) ->P_id_A__U32 x12 <= P_id_A__U32 x11. Hypothesis P_id_A__U61_monotonic : forall x12 x14 x13 x11, (A0 <= x14)/\ (x14 <= x13) -> (A0 <= x12)/\ (x12 <= x11) ->P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Definition marked_measure t := match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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 x11;apply (P_id_A__U43 x11). 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 x11;apply (P_id_MARK x11). 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 x11;apply (P_id_A__ISNATLIST x11). 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 x11;apply (P_id_A__ISNATKIND x11). 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 x11;apply (P_id_A__U31 x12 x11). 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 x11;apply (P_id_A__U53 x11). 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 x11;apply (P_id_A__U41 x13 x12 x11). 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 x11;apply (P_id_A__U22 x11). 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 x11;apply (P_id_A__U51 x13 x12 x11). 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 x11;apply (P_id_A__U42 x12 x11). 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 x11;apply (P_id_A__LENGTH x11). 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 x11;apply (P_id_A__U12 x11). 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 x11;apply (P_id_A__ISNATILISTKIND x11). 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 x11;apply (P_id_A__ISNAT x11). 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 x11;apply (P_id_A__U52 x12 x11). 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 x11;apply (P_id_A__U32 x11). 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 x11;apply (P_id_A__U61 x12 x11). 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 x11;apply (P_id_A__U11 x12 x11). 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 x11;apply (P_id_A__AND x12 x11). 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 x11;apply (P_id_A__U21 x12 x11). 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 x11;apply (P_id_A__ISNATILIST x11). 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_A__ZEROS ). 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_a__zeros => match l with | nil => _ | _::_ => _ end | algebra.F.id_isNatKind => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__isNatIList => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U21 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U41 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_a__and => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U11 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U21 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U61 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U32 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U51 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_isNatIListKind => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_0 => match l with | nil => _ | _::_ => _ end | algebra.F.id_U12 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U52 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__isNat => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U43 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__isNatIListKind => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U12 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U31 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__length => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U42 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U53 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_and => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_cons => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U11 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U51 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_a__U22 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U42 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_length => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_tt => match l with | nil => _ | _::_ => _ end | algebra.F.id_U22 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_s => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U41 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_U52 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_nil => match l with | nil => _ | _::_ => _ end | algebra.F.id_zeros => match l with | nil => _ | _::_ => _ end | algebra.F.id_isNatList => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U53 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U31 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_isNatIList => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__isNatKind => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__isNatList => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U32 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_mark => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U43 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U61 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_isNat => match l with | nil => _ | _::nil => _ | _::_::_ => _ end end. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . 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 . 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 . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . 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 . 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 . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . Qed. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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 |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_isNatKind_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatIList_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U21_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U41_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__and_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U11_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U21_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U61_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U32_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U51_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isNatIListKind_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_U12_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U52_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isNat_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U43_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatIListKind_monotonic; assumption. vm_compute in |-*;intros ;apply P_id_a__U12_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U31_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__length_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U42_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U53_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_and_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_cons_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U11_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U51_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U22_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U42_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_length_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_U22_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_s_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U41_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U52_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_isNatList_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U53_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U31_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isNatIList_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatKind_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatList_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U32_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_mark_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U43_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U61_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isNat_monotonic;assumption. clear f. intros f. case f. vm_compute in |-*;intros ;apply P_id_a__zeros_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNatKind_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatIList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U21_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U41_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__and_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U11_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U21_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U61_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U32_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U51_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNatIListKind_bounded;assumption. vm_compute in |-*;intros ;apply P_id_0_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U12_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U52_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNat_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U43_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatIListKind_bounded; assumption. vm_compute in |-*;intros ;apply P_id_a__U12_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U31_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__length_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U42_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U53_bounded;assumption. vm_compute in |-*;intros ;apply P_id_and_bounded;assumption. vm_compute in |-*;intros ;apply P_id_cons_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U11_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U51_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U22_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U42_bounded;assumption. vm_compute in |-*;intros ;apply P_id_length_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tt_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U22_bounded;assumption. vm_compute in |-*;intros ;apply P_id_s_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U41_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U52_bounded;assumption. vm_compute in |-*;intros ;apply P_id_nil_bounded;assumption. vm_compute in |-*;intros ;apply P_id_zeros_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNatList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U53_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U31_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNatIList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatKind_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNatList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U32_bounded;assumption. vm_compute in |-*;intros ;apply P_id_mark_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U43_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U61_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNat_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_A__U43_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_MARK_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_A__ISNATLIST_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_A__ISNATKIND_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_A__U31_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_A__U53_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_A__U41_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_A__U22_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_A__U51_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_A__U42_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_A__LENGTH_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_A__U12_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_A__ISNATILISTKIND_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_A__ISNAT_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_A__U52_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_A__U32_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_A__U61_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_A__U11_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_A__AND_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_A__U21_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_A__ISNATILIST_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 (Aop.(le_refl)). 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_a__zeros : Z. Hypothesis P_id_isNatKind : Z ->Z. Hypothesis P_id_a__isNatIList : Z ->Z. Hypothesis P_id_a__U21 : Z ->Z ->Z. Hypothesis P_id_U41 : Z ->Z ->Z ->Z. Hypothesis P_id_a__and : Z ->Z ->Z. Hypothesis P_id_a__U11 : Z ->Z ->Z. Hypothesis P_id_U21 : Z ->Z ->Z. Hypothesis P_id_a__U61 : Z ->Z ->Z. Hypothesis P_id_a__U32 : Z ->Z. Hypothesis P_id_U51 : Z ->Z ->Z ->Z. Hypothesis P_id_isNatIListKind : Z ->Z. Hypothesis P_id_0 : Z. Hypothesis P_id_U12 : Z ->Z. Hypothesis P_id_a__U52 : Z ->Z ->Z. Hypothesis P_id_a__isNat : Z ->Z. Hypothesis P_id_U43 : Z ->Z. Hypothesis P_id_a__isNatIListKind : Z ->Z. Hypothesis P_id_a__U12 : Z ->Z. Hypothesis P_id_U31 : Z ->Z ->Z. Hypothesis P_id_a__length : Z ->Z. Hypothesis P_id_a__U42 : Z ->Z ->Z. Hypothesis P_id_U53 : Z ->Z. Hypothesis P_id_and : Z ->Z ->Z. Hypothesis P_id_cons : Z ->Z ->Z. Hypothesis P_id_U11 : Z ->Z ->Z. Hypothesis P_id_a__U51 : Z ->Z ->Z ->Z. Hypothesis P_id_a__U22 : Z ->Z. Hypothesis P_id_U42 : Z ->Z ->Z. Hypothesis P_id_length : Z ->Z. Hypothesis P_id_tt : Z. Hypothesis P_id_U22 : Z ->Z. Hypothesis P_id_s : Z ->Z. Hypothesis P_id_a__U41 : Z ->Z ->Z ->Z. Hypothesis P_id_U52 : Z ->Z ->Z. Hypothesis P_id_nil : Z. Hypothesis P_id_zeros : Z. Hypothesis P_id_isNatList : Z ->Z. Hypothesis P_id_a__U53 : Z ->Z. Hypothesis P_id_a__U31 : Z ->Z ->Z. Hypothesis P_id_isNatIList : Z ->Z. Hypothesis P_id_a__isNatKind : Z ->Z. Hypothesis P_id_a__isNatList : Z ->Z. Hypothesis P_id_U32 : Z ->Z. Hypothesis P_id_mark : Z ->Z. Hypothesis P_id_a__U43 : Z ->Z. Hypothesis P_id_U61 : Z ->Z ->Z. Hypothesis P_id_isNat : Z ->Z. Hypothesis P_id_isNatKind_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) -> P_id_isNatKind x12 <= P_id_isNatKind x11. Hypothesis P_id_a__isNatIList_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Hypothesis P_id_a__U21_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Hypothesis P_id_U41_monotonic : forall x16 x12 x14 x13 x11 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Hypothesis P_id_a__and_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Hypothesis P_id_a__U11_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Hypothesis P_id_U21_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_U21 x12 x14 <= P_id_U21 x11 x13. Hypothesis P_id_a__U61_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Hypothesis P_id_a__U32_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_a__U32 x12 <= P_id_a__U32 x11. Hypothesis P_id_U51_monotonic : forall x16 x12 x14 x13 x11 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Hypothesis P_id_isNatIListKind_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Hypothesis P_id_U12_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Hypothesis P_id_a__U52_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Hypothesis P_id_a__isNat_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) -> P_id_a__isNat x12 <= P_id_a__isNat x11. Hypothesis P_id_U43_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Hypothesis P_id_a__isNatIListKind_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Hypothesis P_id_a__U12_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_a__U12 x12 <= P_id_a__U12 x11. Hypothesis P_id_U31_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_U31 x12 x14 <= P_id_U31 x11 x13. Hypothesis P_id_a__length_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) -> P_id_a__length x12 <= P_id_a__length x11. Hypothesis P_id_a__U42_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Hypothesis P_id_U53_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Hypothesis P_id_and_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_and x12 x14 <= P_id_and x11 x13. Hypothesis P_id_cons_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_cons x12 x14 <= P_id_cons x11 x13. Hypothesis P_id_U11_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_U11 x12 x14 <= P_id_U11 x11 x13. Hypothesis P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Hypothesis P_id_a__U22_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_a__U22 x12 <= P_id_a__U22 x11. Hypothesis P_id_U42_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_U42 x12 x14 <= P_id_U42 x11 x13. Hypothesis P_id_length_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_length x12 <= P_id_length x11. Hypothesis P_id_U22_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Hypothesis P_id_s_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Hypothesis P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Hypothesis P_id_U52_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_U52 x12 x14 <= P_id_U52 x11 x13. Hypothesis P_id_isNatList_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) -> P_id_isNatList x12 <= P_id_isNatList x11. Hypothesis P_id_a__U53_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_a__U53 x12 <= P_id_a__U53 x11. Hypothesis P_id_a__U31_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Hypothesis P_id_isNatIList_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) -> P_id_isNatIList x12 <= P_id_isNatIList x11. Hypothesis P_id_a__isNatKind_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Hypothesis P_id_a__isNatList_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Hypothesis P_id_U32_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Hypothesis P_id_mark_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Hypothesis P_id_a__U43_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_a__U43 x12 <= P_id_a__U43 x11. Hypothesis P_id_U61_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_U61 x12 x14 <= P_id_U61 x11 x13. Hypothesis P_id_isNat_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_isNat x12 <= P_id_isNat x11. Hypothesis P_id_a__zeros_bounded : min_value <= P_id_a__zeros . Hypothesis P_id_isNatKind_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_isNatKind x11. Hypothesis P_id_a__isNatIList_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_a__isNatIList x11. Hypothesis P_id_a__U21_bounded : forall x12 x11, (min_value <= x11) -> (min_value <= x12) ->min_value <= P_id_a__U21 x12 x11. Hypothesis P_id_U41_bounded : forall x12 x13 x11, (min_value <= x11) -> (min_value <= x12) -> (min_value <= x13) ->min_value <= P_id_U41 x13 x12 x11. Hypothesis P_id_a__and_bounded : forall x12 x11, (min_value <= x11) -> (min_value <= x12) ->min_value <= P_id_a__and x12 x11. Hypothesis P_id_a__U11_bounded : forall x12 x11, (min_value <= x11) -> (min_value <= x12) ->min_value <= P_id_a__U11 x12 x11. Hypothesis P_id_U21_bounded : forall x12 x11, (min_value <= x11) ->(min_value <= x12) ->min_value <= P_id_U21 x12 x11. Hypothesis P_id_a__U61_bounded : forall x12 x11, (min_value <= x11) -> (min_value <= x12) ->min_value <= P_id_a__U61 x12 x11. Hypothesis P_id_a__U32_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_a__U32 x11. Hypothesis P_id_U51_bounded : forall x12 x13 x11, (min_value <= x11) -> (min_value <= x12) -> (min_value <= x13) ->min_value <= P_id_U51 x13 x12 x11. Hypothesis P_id_isNatIListKind_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_isNatIListKind x11. Hypothesis P_id_0_bounded : min_value <= P_id_0 . Hypothesis P_id_U12_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_U12 x11. Hypothesis P_id_a__U52_bounded : forall x12 x11, (min_value <= x11) -> (min_value <= x12) ->min_value <= P_id_a__U52 x12 x11. Hypothesis P_id_a__isNat_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_a__isNat x11. Hypothesis P_id_U43_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_U43 x11. Hypothesis P_id_a__isNatIListKind_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_a__isNatIListKind x11. Hypothesis P_id_a__U12_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_a__U12 x11. Hypothesis P_id_U31_bounded : forall x12 x11, (min_value <= x11) ->(min_value <= x12) ->min_value <= P_id_U31 x12 x11. Hypothesis P_id_a__length_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_a__length x11. Hypothesis P_id_a__U42_bounded : forall x12 x11, (min_value <= x11) -> (min_value <= x12) ->min_value <= P_id_a__U42 x12 x11. Hypothesis P_id_U53_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_U53 x11. Hypothesis P_id_and_bounded : forall x12 x11, (min_value <= x11) ->(min_value <= x12) ->min_value <= P_id_and x12 x11. Hypothesis P_id_cons_bounded : forall x12 x11, (min_value <= x11) ->(min_value <= x12) ->min_value <= P_id_cons x12 x11. Hypothesis P_id_U11_bounded : forall x12 x11, (min_value <= x11) ->(min_value <= x12) ->min_value <= P_id_U11 x12 x11. Hypothesis P_id_a__U51_bounded : forall x12 x13 x11, (min_value <= x11) -> (min_value <= x12) -> (min_value <= x13) ->min_value <= P_id_a__U51 x13 x12 x11. Hypothesis P_id_a__U22_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_a__U22 x11. Hypothesis P_id_U42_bounded : forall x12 x11, (min_value <= x11) ->(min_value <= x12) ->min_value <= P_id_U42 x12 x11. Hypothesis P_id_length_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_length x11. Hypothesis P_id_tt_bounded : min_value <= P_id_tt . Hypothesis P_id_U22_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_U22 x11. Hypothesis P_id_s_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_s x11. Hypothesis P_id_a__U41_bounded : forall x12 x13 x11, (min_value <= x11) -> (min_value <= x12) -> (min_value <= x13) ->min_value <= P_id_a__U41 x13 x12 x11. Hypothesis P_id_U52_bounded : forall x12 x11, (min_value <= x11) ->(min_value <= x12) ->min_value <= P_id_U52 x12 x11. Hypothesis P_id_nil_bounded : min_value <= P_id_nil . Hypothesis P_id_zeros_bounded : min_value <= P_id_zeros . Hypothesis P_id_isNatList_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_isNatList x11. Hypothesis P_id_a__U53_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_a__U53 x11. Hypothesis P_id_a__U31_bounded : forall x12 x11, (min_value <= x11) -> (min_value <= x12) ->min_value <= P_id_a__U31 x12 x11. Hypothesis P_id_isNatIList_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_isNatIList x11. Hypothesis P_id_a__isNatKind_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_a__isNatKind x11. Hypothesis P_id_a__isNatList_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_a__isNatList x11. Hypothesis P_id_U32_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_U32 x11. Hypothesis P_id_mark_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_mark x11. Hypothesis P_id_a__U43_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_a__U43 x11. Hypothesis P_id_U61_bounded : forall x12 x11, (min_value <= x11) ->(min_value <= x12) ->min_value <= P_id_U61 x12 x11. Hypothesis P_id_isNat_bounded : forall x11, (min_value <= x11) ->min_value <= P_id_isNat x11. Definition measure := Interp.measure min_value P_id_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12::x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12::x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12::x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12::x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12::x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12::x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12::x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12::x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Hypothesis P_id_A__LENGTH : Z ->Z. Hypothesis P_id_A__U11 : Z ->Z ->Z. Hypothesis P_id_A__U41 : Z ->Z ->Z ->Z. Hypothesis P_id_A__ISNATKIND : Z ->Z. Hypothesis P_id_A__U22 : Z ->Z. Hypothesis P_id_A__U51 : Z ->Z ->Z ->Z. Hypothesis P_id_A__AND : Z ->Z ->Z. Hypothesis P_id_A__ISNATLIST : Z ->Z. Hypothesis P_id_A__U43 : Z ->Z. Hypothesis P_id_A__U31 : Z ->Z ->Z. Hypothesis P_id_A__U53 : Z ->Z. Hypothesis P_id_MARK : Z ->Z. Hypothesis P_id_A__U12 : Z ->Z. Hypothesis P_id_A__U42 : Z ->Z ->Z. Hypothesis P_id_A__ZEROS : Z. Hypothesis P_id_A__ISNAT : Z ->Z. Hypothesis P_id_A__U52 : Z ->Z ->Z. Hypothesis P_id_A__ISNATILISTKIND : Z ->Z. Hypothesis P_id_A__U21 : Z ->Z ->Z. Hypothesis P_id_A__ISNATILIST : Z ->Z. Hypothesis P_id_A__U32 : Z ->Z. Hypothesis P_id_A__U61 : Z ->Z ->Z. Hypothesis P_id_A__LENGTH_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) -> P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Hypothesis P_id_A__U11_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Hypothesis P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Hypothesis P_id_A__ISNATKIND_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Hypothesis P_id_A__U22_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_A__U22 x12 <= P_id_A__U22 x11. Hypothesis P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Hypothesis P_id_A__AND_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Hypothesis P_id_A__ISNATLIST_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Hypothesis P_id_A__U43_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_A__U43 x12 <= P_id_A__U43 x11. Hypothesis P_id_A__U31_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Hypothesis P_id_A__U53_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_A__U53 x12 <= P_id_A__U53 x11. Hypothesis P_id_MARK_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Hypothesis P_id_A__U12_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_A__U12 x12 <= P_id_A__U12 x11. Hypothesis P_id_A__U42_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Hypothesis P_id_A__ISNAT_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) -> P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Hypothesis P_id_A__U52_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Hypothesis P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Hypothesis P_id_A__U21_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Hypothesis P_id_A__ISNATILIST_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Hypothesis P_id_A__U32_monotonic : forall x12 x11, (min_value <= x12)/\ (x12 <= x11) ->P_id_A__U32 x12 <= P_id_A__U32 x11. Hypothesis P_id_A__U61_monotonic : forall x12 x14 x13 x11, (min_value <= x14)/\ (x14 <= x13) -> (min_value <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Definition marked_measure := Interp.marked_measure min_value P_id_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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 := (* *) | DP_R_xml_0_0 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_a__isNatList (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) (* *) | DP_R_xml_0_1 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatList (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) (* *) | DP_R_xml_0_2 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) (* *) | DP_R_xml_0_3 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) (* *) | DP_R_xml_0_4 : forall x12 x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_a__isNatList (x2::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) (* *) | DP_R_xml_0_5 : forall x12 x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatList (x2::nil)) (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) (* *) | DP_R_xml_0_6 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_7 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_8 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_a__isNatIList (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) (* *) | DP_R_xml_0_9 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatIList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) (* *) | DP_R_xml_0_10 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_11 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_12 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_a__isNatList (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) (* *) | DP_R_xml_0_13 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) (* *) | DP_R_xml_0_14 : forall x4 x12 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Term algebra.F.id_mark (x4::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) (* *) | DP_R_xml_0_15 : forall x4 x12 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) (* *) | DP_R_xml_0_16 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) (* *) | DP_R_xml_0_17 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_18 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_19 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_20 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_21 : forall x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x2::nil))::x2::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_22 : forall x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x2::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_23 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1:: x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_24 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_25 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_26 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_27 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_28 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_29 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_30 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1:: x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_31 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_32 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_33 : forall x4 x6 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x6::x4::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatList (x4::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x4::nil))::nil))::(algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isNat (x6::nil)):: (algebra.Alg.Term algebra.F.id_isNatKind (x6::nil))::nil))::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) (* *) | DP_R_xml_0_34 : forall x4 x6 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x6::x4::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatList (x4::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x4::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isNat (x6::nil))::(algebra.Alg.Term algebra.F.id_isNatKind (x6::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) (* *) | DP_R_xml_0_35 : forall x4 x6 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x6::x4::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatList (x4::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x4::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) (* *) | DP_R_xml_0_36 : forall x4 x6 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x6::x4::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatList (x4::nil)) (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) (* *) | DP_R_xml_0_37 : forall x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_zeros nil) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__zeros nil) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_38 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x7::x8::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_39 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x7::x8::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_40 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_mark (x5::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_41 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_42 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatList (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_43 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_44 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_45 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_mark (x5::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_46 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_47 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNat (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNat (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_48 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x7::x8::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_49 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x7::x8::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_50 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_mark (x5::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_51 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_52 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x7::x8::x9::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_53 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x7::x8::x9::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_54 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x7::x8::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_55 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x7::x8::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_56 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_mark (x5::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_57 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_58 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIList (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatIList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_59 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_60 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_61 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_62 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_63 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_mark (x5::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_64 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_65 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U61 (x7::x8::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_66 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U61 (x7::x8::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_67 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Term algebra.F.id_mark (x5::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_68 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_69 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_70 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_71 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_72 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatKind (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNatKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_73 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x7::x8::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_74 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module ddp := dp.MakeDP(algebra.EQT). Lemma R_xml_0_dp_step_spec : forall x y, (ddp.dp_step R_xml_0_deep_rew.R_xml_0_rules x y) -> exists f, exists l1, exists l2, y = algebra.Alg.Term f l2/\ (closure.refl_trans_clos (closure.one_step_list (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) ) l1 l2)/\ (ddp.dp R_xml_0_deep_rew.R_xml_0_rules x (algebra.Alg.Term f l1)). Proof. intros x y H. induction H. inversion H. subst. destruct t0. refine ((False_ind) _ _). refine (R_xml_0_deep_rew.R_xml_0_non_var H0). simpl in H|-*. exists a. exists ((List.map) (algebra.Alg.apply_subst sigma) l). exists ((List.map) (algebra.Alg.apply_subst sigma) l). repeat (constructor). assumption. exists f. exists l2. exists l1. constructor. constructor. constructor. constructor. rewrite <- closure.rwr_list_trans_clos_one_step_list. assumption. assumption. Qed. Ltac included_dp_tac H := injection H;clear H;intros;subst; repeat (match goal with | H: closure.refl_trans_clos (closure.one_step_list _) (_::_) _ |- _=> let x := fresh "x" in let l := fresh "l" in let h1 := fresh "h" in let h2 := fresh "h" in let h3 := fresh "h" in destruct (@algebra.EQT_ext.one_step_list_star_decompose_cons _ _ _ _ H) as [x [l[h1[h2 h3]]]];clear H;subst | H: closure.refl_trans_clos (closure.one_step_list _) nil _ |- _ => rewrite (@algebra.EQT_ext.one_step_list_star_decompose_nil _ _ H) in *;clear H end );simpl; econstructor eassumption . Ltac dp_concl_tac h2 h cont_tac t := match t with | False => let h' := fresh "a" in (set (h':=t) in *;cont_tac h'; repeat ( let e := type of h in (match e with | ?t => unfold t in h|-; (case h; [abstract (clear h;intros h;injection h; clear h;intros ;subst; included_dp_tac h2)| clear h;intros h;clear t]) | ?t => unfold t in h|-;elim h end ) )) | or ?a ?b => let cont_tac h' := let h'' := fresh "a" in (set (h'':=or a h') in *;cont_tac h'') in (dp_concl_tac h2 h cont_tac b) end . Module WF_DP_R_xml_0. Inductive DP_R_xml_0_non_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_1_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_non_scc_1 (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_mark (x5::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Lemma acc_DP_R_xml_0_non_scc_1 : forall x y, (DP_R_xml_0_non_scc_1 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_2 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_2_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x5::nil)) x11) -> DP_R_xml_0_non_scc_2 (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_mark (x5::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Lemma acc_DP_R_xml_0_non_scc_2 : forall x y, (DP_R_xml_0_non_scc_2 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_3 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_3_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x5::nil)) x11) -> DP_R_xml_0_non_scc_3 (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_mark (x5::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Lemma acc_DP_R_xml_0_non_scc_3 : forall x y, (DP_R_xml_0_non_scc_3 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_4 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_4_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_non_scc_4 (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_mark (x5::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Lemma acc_DP_R_xml_0_non_scc_4 : forall x y, (DP_R_xml_0_non_scc_4 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_5 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_5_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x5::nil)) x11) -> DP_R_xml_0_non_scc_5 (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_mark (x5::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Lemma acc_DP_R_xml_0_non_scc_5 : forall x y, (DP_R_xml_0_non_scc_5 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_6 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_6_0 : forall x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_zeros nil) x11) -> DP_R_xml_0_non_scc_6 (algebra.Alg.Term algebra.F.id_a__zeros nil) (algebra.Alg.Term algebra.F.id_mark (x11::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; (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 := (* *) | DP_R_xml_0_non_scc_7_0 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_non_scc_7 (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_a__isNatList (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::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; (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 := (* *) | DP_R_xml_0_non_scc_8_0 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_non_scc_8 (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_a__isNatIList (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::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; (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 := (* *) | DP_R_xml_0_non_scc_9_0 : forall x12 x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_non_scc_9 (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_a__isNatList (x2::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::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; (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 := (* *) | DP_R_xml_0_non_scc_10_0 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_non_scc_10 (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::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; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_11 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_11_0 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_non_scc_11 (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_a__isNatList (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) . Lemma acc_DP_R_xml_0_non_scc_11 : forall x y, (DP_R_xml_0_non_scc_11 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (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_12 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_0 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatList (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_1 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1:: x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_2 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_3 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_4 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_5 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_6 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_7 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_8 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x5::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_9 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatList (x5::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_10 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_11 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_12 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_13 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_14 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_15 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_16 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_17 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_18 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_19 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_20 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_21 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_22 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_23 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNat (x5::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNat (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_24 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_25 : forall x12 x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatList (x2::nil)) (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_26 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_27 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x5::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_28 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_29 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_30 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatIList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_31 : forall x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x2::nil))::x2::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_32 : forall x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x2::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_33 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1:: x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_34 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_35 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_36 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_37 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_38 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_39 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_40 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x5::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_41 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIList (x5::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatIList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_42 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_43 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_44 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_45 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_46 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_47 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_48 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U61 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_49 : forall x4 x12 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Term algebra.F.id_mark (x4::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_50 : forall x4 x6 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x6::x4::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatList (x4::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x4::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isNat (x6::nil)):: (algebra.Alg.Term algebra.F.id_isNatKind (x6::nil))::nil))::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) (* *) | DP_R_xml_0_scc_12_51 : forall x4 x12 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_52 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U61 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_53 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x5::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Term algebra.F.id_mark (x5::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_54 : forall x4 x6 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x6::x4::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatList (x4::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x4::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isNat (x6::nil)):: (algebra.Alg.Term algebra.F.id_isNatKind (x6::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) (* *) | DP_R_xml_0_scc_12_55 : forall x4 x6 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x6::x4::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatList (x4::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x4::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) (* *) | DP_R_xml_0_scc_12_56 : forall x4 x6 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x6::x4::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatList (x4::nil)) (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) (* *) | DP_R_xml_0_scc_12_57 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x5::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_58 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_59 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_60 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_61 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatKind (x5::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_a__isNatKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_62 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_63 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12. Inductive DP_R_xml_0_scc_12_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_0 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNatList (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_1 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1::x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_2 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_3 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNatList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_4 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_5 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_6 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_7 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_8 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_9 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatList (x5::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNatList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_10 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_11 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_12 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_13 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_14 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_15 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_16 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_17 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)):: x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_18 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_19 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_20 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_21 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_22 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_23 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNat (x5::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNat (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_24 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_25 : forall x12 x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNatList (x2::nil)) (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_26 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_27 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_28 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_29 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_30 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNatIList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_31 : forall x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x2::nil)):: x2::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_32 : forall x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x2::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_33 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1::x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_34 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_35 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_36 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_37 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_38 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_39 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_40 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_41 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIList (x5::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNatIList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_42 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_43 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_44 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_45 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_46 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_47 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_48 : forall x4 x12 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Term algebra.F.id_mark (x4::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_49 : forall x4 x6 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x6::x4::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatList (x4::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x4::nil))::nil))::(algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isNat (x6::nil)):: (algebra.Alg.Term algebra.F.id_isNatKind (x6::nil))::nil))::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_50 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_51 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_52 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_53 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_54 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_55 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Inductive DP_R_xml_0_scc_12_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_strict_0 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U61 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_strict (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_strict_1 : forall x4 x12 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x11) -> DP_R_xml_0_scc_12_strict (algebra.Alg.Term algebra.F.id_mark (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_strict_2 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U61 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_strict (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_strict_3 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x5::nil)) x11) -> DP_R_xml_0_scc_12_strict (algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Term algebra.F.id_mark (x5::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_strict_4 : forall x4 x6 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x6::x4::nil)) x11) -> DP_R_xml_0_scc_12_strict (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatList (x4::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x4::nil))::nil))::(algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isNat (x6::nil)):: (algebra.Alg.Term algebra.F.id_isNatKind (x6::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) (* *) | DP_R_xml_0_scc_12_strict_5 : forall x4 x6 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x6::x4::nil)) x11) -> DP_R_xml_0_scc_12_strict (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatList (x4::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x4::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) (* *) | DP_R_xml_0_scc_12_strict_6 : forall x4 x6 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x6::x4::nil)) x11) -> DP_R_xml_0_scc_12_strict (algebra.Alg.Term algebra.F.id_a__isNatList (x4::nil)) (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) (* *) | DP_R_xml_0_scc_12_strict_7 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x5::nil)) x11) -> DP_R_xml_0_scc_12_strict (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large. Inductive DP_R_xml_0_scc_12_large_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_1_0 : forall x4 x6 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x6::x4::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_1 (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatList (x4::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x4::nil))::nil))::(algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isNat (x6::nil)):: (algebra.Alg.Term algebra.F.id_isNatKind (x6::nil))::nil))::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_1_1 : forall x4 x12 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x11) -> DP_R_xml_0_scc_12_large_scc_1 (algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Term algebra.F.id_mark (x4::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_1. Inductive DP_R_xml_0_scc_12_large_scc_1_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_1_large_0 : forall x4 x6 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x6::x4::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatList (x4::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x4::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isNat (x6::nil))::(algebra.Alg.Term algebra.F.id_isNatKind (x6::nil))::nil))::nil)):: x4::nil)) (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) . Inductive DP_R_xml_0_scc_12_large_scc_1_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_1_strict_0 : forall x4 x12 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x11) -> DP_R_xml_0_scc_12_large_scc_1_strict (algebra.Alg.Term algebra.F.id_a__length ((algebra.Alg.Term algebra.F.id_mark (x4::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_1_large. Inductive DP_R_xml_0_scc_12_large_scc_1_large_non_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_1_large_non_scc_1_0 : forall x4 x6 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x6::x4::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_1_large_non_scc_1 (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatList (x4::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x4::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isNat (x6::nil)):: (algebra.Alg.Term algebra.F.id_isNatKind (x6::nil))::nil))::nil)):: x4::nil)) (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_1_large_non_scc_1 : forall x y, (DP_R_xml_0_scc_12_large_scc_1_large_non_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_1.DP_R_xml_0_scc_12_large_scc_1_large x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_1.DP_R_xml_0_scc_12_large_scc_1_large . Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_12_large_scc_1_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_1_large_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||(fail))). Qed. End WF_DP_R_xml_0_scc_12_large_scc_1_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_a__zeros := 1. Definition P_id_isNatKind (x11:Z) := 0. Definition P_id_a__isNatIList (x11:Z) := 2 + 1* x11. Definition P_id_a__U21 (x11:Z) (x12:Z) := 1. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 1 + 1* x11 + 1* x13. Definition P_id_a__and (x11:Z) (x12:Z) := 1* x11 + 2* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 1* x11. Definition P_id_U21 (x11:Z) (x12:Z) := 0. Definition P_id_a__U61 (x11:Z) (x12:Z) := 0. Definition P_id_a__U32 (x11:Z) := 1* x11. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 3* x13. Definition P_id_isNatIListKind (x11:Z) := 0. Definition P_id_0 := 0. Definition P_id_U12 (x11:Z) := 0. Definition P_id_a__U52 (x11:Z) (x12:Z) := 1* x12. Definition P_id_a__isNat (x11:Z) := 1. Definition P_id_U43 (x11:Z) := 1* x11. Definition P_id_a__isNatIListKind (x11:Z) := 1. Definition P_id_a__U12 (x11:Z) := 1. Definition P_id_U31 (x11:Z) (x12:Z) := 1* x12. Definition P_id_a__length (x11:Z) := 0. Definition P_id_a__U42 (x11:Z) (x12:Z) := 1 + 1* x11 + 1* x12. Definition P_id_U53 (x11:Z) := 1* x11. Definition P_id_and (x11:Z) (x12:Z) := 1* x11 + 2* x12. Definition P_id_cons (x11:Z) (x12:Z) := 3* x12. Definition P_id_U11 (x11:Z) (x12:Z) := 1* x11. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 3* x13. Definition P_id_a__U22 (x11:Z) := 1* x11. Definition P_id_U42 (x11:Z) (x12:Z) := 1 + 1* x11 + 1* x12. Definition P_id_length (x11:Z) := 0. Definition P_id_tt := 1. Definition P_id_U22 (x11:Z) := 1* x11. Definition P_id_s (x11:Z) := 0. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 1 + 1* x11 + 1* x13. Definition P_id_U52 (x11:Z) (x12:Z) := 1* x12. Definition P_id_nil := 1. Definition P_id_zeros := 0. Definition P_id_isNatList (x11:Z) := 1* x11. Definition P_id_a__U53 (x11:Z) := 1* x11. Definition P_id_a__U31 (x11:Z) (x12:Z) := 1* x12. Definition P_id_isNatIList (x11:Z) := 1 + 1* x11. Definition P_id_a__isNatKind (x11:Z) := 1. Definition P_id_a__isNatList (x11:Z) := 1* x11. Definition P_id_U32 (x11:Z) := 1* x11. Definition P_id_mark (x11:Z) := 1 + 1* x11. Definition P_id_a__U43 (x11:Z) := 1* x11. Definition P_id_U61 (x11:Z) (x12:Z) := 0. Definition P_id_isNat (x11:Z) := 0. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13::x12:: x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12::x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13::x12:: x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12::x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12::x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12::x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12::x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12::x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12::x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12::x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 1* x11. Definition P_id_A__U11 (x11:Z) (x12:Z) := 0. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 0. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__AND (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATLIST (x11:Z) := 0. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 0. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 0. Definition P_id_A__U52 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILISTKIND (x11:Z) := 0. Definition P_id_A__U21 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 2* x11 + 1* x12. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13:: x12::x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13:: x12::x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_1_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_1_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_12_large_scc_1_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_1_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_12_large_scc_1_large := WF_DP_R_xml_0_scc_12_large_scc_1_large.wf. Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large.DP_R_xml_0_scc_12_large_scc_1. Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_1_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx';apply DP_R_xml_0_scc_12_large_scc_1_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_12_large_scc_1_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_12_large_scc_1_large. Qed. End WF_DP_R_xml_0_scc_12_large_scc_1. Definition wf_DP_R_xml_0_scc_12_large_scc_1 := WF_DP_R_xml_0_scc_12_large_scc_1.wf. Lemma acc_DP_R_xml_0_scc_12_large_scc_1 : forall x y, (DP_R_xml_0_scc_12_large_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_12.DP_R_xml_0_scc_12_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_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_12_large_scc_1. Qed. Inductive DP_R_xml_0_scc_12_large_scc_2 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_0 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1::x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_1 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)):: x3::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_2 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_3 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_4 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_5 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_6 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatList (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_7 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_8 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_9 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_10 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_11 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_12 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_13 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_14 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatList (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_15 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_16 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_17 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_18 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_19 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_20 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_21 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_22 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_23 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNat (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNat (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_24 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_25 : forall x12 x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatList (x2::nil)) (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_26 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_27 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_28 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_29 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)):: x3::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_30 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatIList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_31 : forall x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x2::nil))::x2::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_32 : forall x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x2::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_33 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1::x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_34 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_35 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_36 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_37 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_38 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_39 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_40 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_41 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIList (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatIList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_42 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_43 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_44 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_45 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_46 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_47 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_48 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_49 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_50 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_51 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_52 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_53 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2. Inductive DP_R_xml_0_scc_12_large_scc_2_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_0 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1:: x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_1 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_2 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNatList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_3 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_4 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_5 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_6 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNatList (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_7 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_8 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_9 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_10 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_11 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_12 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_13 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_14 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatList (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNatList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_15 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_16 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_17 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_18 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_19 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_20 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_21 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_22 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_23 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNat (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNat (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_24 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_25 : forall x12 x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNatList (x2::nil)) (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_26 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_27 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_28 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_29 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_30 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNatIList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_31 : forall x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x2::nil))::x2::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_32 : forall x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x2::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_33 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1:: x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_34 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_35 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_36 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_37 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_38 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_39 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_40 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_41 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIList (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNatIList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_42 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_43 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_44 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_45 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_46 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_47 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_48 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_49 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_50 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_51 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_52 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Inductive DP_R_xml_0_scc_12_large_scc_2_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_strict_0 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_strict (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_0 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1:: x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_1 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_2 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNatList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_3 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_4 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_5 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_6 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNatList (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_7 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_8 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_9 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_10 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_11 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_12 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_13 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_14 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatList (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNatList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_15 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_16 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_17 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_18 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_19 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_20 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_21 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_22 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_23 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNat (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNat (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_24 : forall x12 x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNatList (x2::nil)) (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_25 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_26 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNatIList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_27 : forall x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x2::nil))::x2::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_28 : forall x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x2::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_29 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1:: x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_30 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_31 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_32 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_33 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_34 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8:: x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_35 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_36 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_37 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_38 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_39 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_40 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_41 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_42 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_43 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_44 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Inductive DP_R_xml_0_scc_12_large_scc_2_large_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_strict_0 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_strict (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_strict_1 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_strict (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_strict_2 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_strict (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_strict_3 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_strict (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8:: x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_strict_4 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_strict (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_strict_5 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_strict (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_strict_6 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_strict (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_strict_7 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIList (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_strict (algebra.Alg.Term algebra.F.id_a__isNatIList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_0 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)):: x3::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_1 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_2 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)):: x1::x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_3 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_4 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)):: x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_5 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatList (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_6 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_7 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_8 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_9 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_10 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_11 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatList (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_12 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_13 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_14 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_15 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_16 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_17 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_18 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_19 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_20 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_21 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_22 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_23 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_24 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNat (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNat (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_25 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_26 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8:: x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_27 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_28 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_29 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_30 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_31 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_32 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_33 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_34 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_35 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_0 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_1 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNatList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_2 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__U51 (( algebra.Alg.Term algebra.F.id_a__and (( algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: ( algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)):: x1:: x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_3 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_4 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__U11 (( algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)):: x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_5 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNatList (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_6 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__and (( algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: ( algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_7 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_8 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_9 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatList (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNatList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_10 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_11 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_12 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__and (( algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: ( algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_13 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_14 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_15 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__U21 (( algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_16 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_17 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_18 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__U21 (( algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_19 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_20 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_21 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_22 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNat (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNat (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_23 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_24 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__U51 (( algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8:: x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_25 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_26 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__U52 (( algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_27 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_28 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_29 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__and (( algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_30 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_31 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_32 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_33 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_strict_0 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_strict (algebra.Alg.Term algebra.F.id_a__U11 (( algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_strict_1 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_strict (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_0 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_1 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNatList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_2 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1::x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_3 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_4 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_5 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNatList (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_6 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_7 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_8 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_9 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatList (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNatList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_10 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_11 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_12 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_13 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_14 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_15 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_16 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_17 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_18 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_19 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_20 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_21 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_22 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNat (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNat (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_23 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_24 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_25 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_26 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_27 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_28 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_29 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_30 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_31 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_32 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_strict_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_strict (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_0 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_1 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__isNatList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_2 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1::x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_3 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_4 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__isNatList (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_5 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_6 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_7 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatList (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__isNatList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_8 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_9 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_10 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_11 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_12 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_13 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_14 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_15 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_16 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_17 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_18 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNat (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__isNat (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_19 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_20 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_21 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_22 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_23 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_24 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_25 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_26 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_strict_0 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_strict (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_strict_1 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_strict (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_strict_2 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_strict (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_strict_3 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_strict (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_strict_4 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_strict (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_strict_5 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_strict (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_0 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_1 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1:: x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_2 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_0 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1:: x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_1 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) . Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_strict_0 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_strict (algebra.Alg.Term algebra.F.id_a__isNatList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_non_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_non_scc_1_0 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_non_scc_1 (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)):: x3::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_non_scc_1 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_non_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_non_scc_2 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_non_scc_2_0 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_non_scc_2 (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1:: x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_non_scc_2 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_non_scc_2 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large 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_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large . Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_non_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (fail)))). Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_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_a__zeros := 1. Definition P_id_isNatKind (x11:Z) := 0. Definition P_id_a__isNatIList (x11:Z) := 3 + 2* x11. Definition P_id_a__U21 (x11:Z) (x12:Z) := 1* x11. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 3 + 2* x13. Definition P_id_a__and (x11:Z) (x12:Z) := 1 + 2* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 1* x11. Definition P_id_U21 (x11:Z) (x12:Z) := 1* x11. Definition P_id_a__U61 (x11:Z) (x12:Z) := 1* x11. Definition P_id_a__U32 (x11:Z) := 1. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 1. Definition P_id_isNatIListKind (x11:Z) := 0. Definition P_id_0 := 1. Definition P_id_U12 (x11:Z) := 0. Definition P_id_a__U52 (x11:Z) (x12:Z) := 1. Definition P_id_a__isNat (x11:Z) := 1* x11. Definition P_id_U43 (x11:Z) := 1. Definition P_id_a__isNatIListKind (x11:Z) := 1. Definition P_id_a__U12 (x11:Z) := 1. Definition P_id_U31 (x11:Z) (x12:Z) := 0. Definition P_id_a__length (x11:Z) := 1. Definition P_id_a__U42 (x11:Z) (x12:Z) := 1. Definition P_id_U53 (x11:Z) := 0. Definition P_id_and (x11:Z) (x12:Z) := 2* x12. Definition P_id_cons (x11:Z) (x12:Z) := 1* x11 + 1* x12. Definition P_id_U11 (x11:Z) (x12:Z) := 1* x11. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 1. Definition P_id_a__U22 (x11:Z) := 1. Definition P_id_U42 (x11:Z) (x12:Z) := 0. Definition P_id_length (x11:Z) := 1. Definition P_id_tt := 1. Definition P_id_U22 (x11:Z) := 0. Definition P_id_s (x11:Z) := 1. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 3 + 2* x13. Definition P_id_U52 (x11:Z) (x12:Z) := 0. Definition P_id_nil := 0. Definition P_id_zeros := 0. Definition P_id_isNatList (x11:Z) := 1. Definition P_id_a__U53 (x11:Z) := 1. Definition P_id_a__U31 (x11:Z) (x12:Z) := 1. Definition P_id_isNatIList (x11:Z) := 3 + 2* x11. Definition P_id_a__isNatKind (x11:Z) := 1. Definition P_id_a__isNatList (x11:Z) := 1. Definition P_id_U32 (x11:Z) := 1. Definition P_id_mark (x11:Z) := 1 + 1* x11. Definition P_id_a__U43 (x11:Z) := 1. Definition P_id_U61 (x11:Z) (x12:Z) := 1* x11. Definition P_id_isNat (x11:Z) := 1* x11. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13::x12:: x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12:: x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13::x12:: x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12:: x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12:: x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12:: x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12:: x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13:: x12::x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12:: x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13:: x12::x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12:: x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12:: x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 0. Definition P_id_A__U11 (x11:Z) (x12:Z) := 0. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 0. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 1* x12 + 1* x13 . Definition P_id_A__AND (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATLIST (x11:Z) := 1* x11. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 0. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 0. Definition P_id_A__U52 (x11:Z) (x12:Z) := 1* x11 + 1* x12. Definition P_id_A__ISNATILISTKIND (x11:Z) := 0. Definition P_id_A__U21 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 0. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large.wf . Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1 . Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large) . clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1_large. Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1. Definition wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1 := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1.wf . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_1. Qed. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_2 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_2_0 : forall x8 x9 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x7::x8::x9::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_2 (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_2 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_2 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large 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_12_large_scc_2_large_large_scc_1_large_large_large_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_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_3 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_3_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatList (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_3 (algebra.Alg.Term algebra.F.id_a__isNatList (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_3 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_3 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large 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_12_large_scc_2_large_large_scc_1_large_large_large_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_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_4 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_4_0 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_4 (algebra.Alg.Term algebra.F.id_a__isNatList (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_4 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_4 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large 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_12_large_scc_2_large_large_scc_1_large_large_large_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_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_5 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_5_0 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_5 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_5 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_5 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large 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_12_large_scc_2_large_large_scc_1_large_large_large_non_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_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_6 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_6_0 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_6 (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_6 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_6 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large 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_12_large_scc_2_large_large_scc_1_large_large_large_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_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_1 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_2 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_3 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_4 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_5 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_6 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_7 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_8 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_9 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_10 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNat (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_a__isNat (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_11 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_12 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_13 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_14 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_15 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_16 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_17 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_a__isNatKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_18 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_1 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_2 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_3 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_4 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_5 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_6 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_7 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_8 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_9 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_10 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_11 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_12 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_13 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_14 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_15 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_strict_0 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_strict (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_strict_1 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_strict (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_strict_2 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNat (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_strict (algebra.Alg.Term algebra.F.id_a__isNat (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_1 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_2 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_3 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_4 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_5 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_6 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_7 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_8 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_9 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_10 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_11 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_12 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_1 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_2 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil))::x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_3 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_4 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_5 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_6 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_7 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_8 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_9 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_10 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_11 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_strict_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_strict (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_1 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_2 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_3 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_4 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_5 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_6 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_7 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_8 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_9 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_a__isNatKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_10 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_strict_0 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_length (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_strict (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_1_0 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_1 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_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_a__zeros := 1. Definition P_id_isNatKind (x11:Z) := 0. Definition P_id_a__isNatIList (x11:Z) := 3. Definition P_id_a__U21 (x11:Z) (x12:Z) := 2. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 1. Definition P_id_a__and (x11:Z) (x12:Z) := 1* x11 + 2* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 1* x11. Definition P_id_U21 (x11:Z) (x12:Z) := 0. Definition P_id_a__U61 (x11:Z) (x12:Z) := 1 + 1* x11 + 1* x12. Definition P_id_a__U32 (x11:Z) := 2. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 3* x13. Definition P_id_isNatIListKind (x11:Z) := 0. Definition P_id_0 := 0. Definition P_id_U12 (x11:Z) := 2. Definition P_id_a__U52 (x11:Z) (x12:Z) := 1 + 3* x12. Definition P_id_a__isNat (x11:Z) := 2. Definition P_id_U43 (x11:Z) := 3. Definition P_id_a__isNatIListKind (x11:Z) := 2. Definition P_id_a__U12 (x11:Z) := 2. Definition P_id_U31 (x11:Z) (x12:Z) := 1. Definition P_id_a__length (x11:Z) := 1* x11. Definition P_id_a__U42 (x11:Z) (x12:Z) := 3. Definition P_id_U53 (x11:Z) := 1* x11. Definition P_id_and (x11:Z) (x12:Z) := 1* x11 + 2* x12. Definition P_id_cons (x11:Z) (x12:Z) := 1 + 3* x12. Definition P_id_U11 (x11:Z) (x12:Z) := 1* x11. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 1 + 3* x13 . Definition P_id_a__U22 (x11:Z) := 2. Definition P_id_U42 (x11:Z) (x12:Z) := 3. Definition P_id_length (x11:Z) := 1* x11. Definition P_id_tt := 2. Definition P_id_U22 (x11:Z) := 0. Definition P_id_s (x11:Z) := 1 + 1* x11. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 3. Definition P_id_U52 (x11:Z) (x12:Z) := 1 + 3* x12. Definition P_id_nil := 2. Definition P_id_zeros := 0. Definition P_id_isNatList (x11:Z) := 2* x11. Definition P_id_a__U53 (x11:Z) := 1* x11. Definition P_id_a__U31 (x11:Z) (x12:Z) := 3. Definition P_id_isNatIList (x11:Z) := 1. Definition P_id_a__isNatKind (x11:Z) := 2. Definition P_id_a__isNatList (x11:Z) := 2* x11. Definition P_id_U32 (x11:Z) := 2. Definition P_id_mark (x11:Z) := 2 + 1* x11. Definition P_id_a__U43 (x11:Z) := 3. Definition P_id_U61 (x11:Z) (x12:Z) := 1 + 1* x11 + 1* x12. Definition P_id_isNat (x11:Z) := 0. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13:: x12::x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12:: x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13:: x12::x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12:: x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12:: x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12:: x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12:: x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12:: x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12:: x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12:: x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 0. Definition P_id_A__U11 (x11:Z) (x12:Z) := 0. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 1* x11. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__AND (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATLIST (x11:Z) := 0. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 0. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 0. Definition P_id_A__U52 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILISTKIND (x11:Z) := 0. Definition P_id_A__U21 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 0. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_1. Definition wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_1 := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_1.wf . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_1 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_1. Qed. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_non_scc_2 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_non_scc_2_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_non_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_non_scc_2 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_non_scc_2 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large 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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_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_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_non_scc_3 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_non_scc_3_0 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_non_scc_3 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_non_scc_3 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_non_scc_3 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large 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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_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_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_1 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_2 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_3 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_4 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4 (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_5 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_6 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_7 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_1 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_2 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_3 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_4 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_5 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_6 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_strict_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_strict (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_1 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_2 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_strict_0 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_strict (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x7::nil)):: x8::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_strict_1 : forall x8 x11 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x7::x8::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_strict (algebra.Alg.Term algebra.F.id_mark (x7::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_strict_2 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNatIListKind (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_strict (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_strict_3 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_strict (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_1 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_large_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_large (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_strict_0 : forall x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x5::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_strict (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_mark (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_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_a__zeros := 1. Definition P_id_isNatKind (x11:Z) := 0. Definition P_id_a__isNatIList (x11:Z) := 1. Definition P_id_a__U21 (x11:Z) (x12:Z) := 3 + 2* x12. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_a__and (x11:Z) (x12:Z) := 1* x11 + 2* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 1. Definition P_id_U21 (x11:Z) (x12:Z) := 1 + 2* x12. Definition P_id_a__U61 (x11:Z) (x12:Z) := 1 + 1* x11 + 2* x12. Definition P_id_a__U32 (x11:Z) := 1. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 1* x12 + 3* x13. Definition P_id_isNatIListKind (x11:Z) := 0. Definition P_id_0 := 0. Definition P_id_U12 (x11:Z) := 1. Definition P_id_a__U52 (x11:Z) (x12:Z) := 1* x12. Definition P_id_a__isNat (x11:Z) := 1 + 2* x11. Definition P_id_U43 (x11:Z) := 0. Definition P_id_a__isNatIListKind (x11:Z) := 1. Definition P_id_a__U12 (x11:Z) := 1. Definition P_id_U31 (x11:Z) (x12:Z) := 1* x11. Definition P_id_a__length (x11:Z) := 1* x11. Definition P_id_a__U42 (x11:Z) (x12:Z) := 1. Definition P_id_U53 (x11:Z) := 1* x11. Definition P_id_and (x11:Z) (x12:Z) := 1* x11 + 2* x12. Definition P_id_cons (x11:Z) (x12:Z) := 1 + 2* x11 + 3* x12. Definition P_id_U11 (x11:Z) (x12:Z) := 0. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 2* x12 + 3* x13. Definition P_id_a__U22 (x11:Z) := 2 + 1* x11. Definition P_id_U42 (x11:Z) (x12:Z) := 0. Definition P_id_length (x11:Z) := 1* x11. Definition P_id_tt := 1. Definition P_id_U22 (x11:Z) := 2 + 1* x11. Definition P_id_s (x11:Z) := 1 + 1* x11. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 1. Definition P_id_U52 (x11:Z) (x12:Z) := 1* x12. Definition P_id_nil := 2. Definition P_id_zeros := 0. Definition P_id_isNatList (x11:Z) := 1* x11. Definition P_id_a__U53 (x11:Z) := 1* x11. Definition P_id_a__U31 (x11:Z) (x12:Z) := 1* x11. Definition P_id_isNatIList (x11:Z) := 0. Definition P_id_a__isNatKind (x11:Z) := 1. Definition P_id_a__isNatList (x11:Z) := 1* x11. Definition P_id_U32 (x11:Z) := 1. Definition P_id_mark (x11:Z) := 1 + 2* x11. Definition P_id_a__U43 (x11:Z) := 1. Definition P_id_U61 (x11:Z) (x12:Z) := 1 + 1* x11 + 2* x12. Definition P_id_isNat (x11:Z) := 1* x11. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) -> (0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) -> (0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12::x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12::x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12::x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12::x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12::x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12::x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12::x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12::x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic; assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded; assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 0. Definition P_id_A__U11 (x11:Z) (x12:Z) := 0. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 0. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__AND (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATLIST (x11:Z) := 0. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 2* x11. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 0. Definition P_id_A__U52 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILISTKIND (x11:Z) := 0. Definition P_id_A__U21 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 0. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13:: x12::x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13:: x12::x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic; assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded; assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic; assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_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_a__zeros := 1. Definition P_id_isNatKind (x11:Z) := 0. Definition P_id_a__isNatIList (x11:Z) := 3 + 3* x11. Definition P_id_a__U21 (x11:Z) (x12:Z) := 1. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 2 + 2* x11 + 2* x13. Definition P_id_a__and (x11:Z) (x12:Z) := 1* x11 + 2* x12 . Definition P_id_a__U11 (x11:Z) (x12:Z) := 1. Definition P_id_U21 (x11:Z) (x12:Z) := 0. Definition P_id_a__U61 (x11:Z) (x12:Z) := 1 + 1* x11 + 2* x12. Definition P_id_a__U32 (x11:Z) := 1. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 1* x13. Definition P_id_isNatIListKind (x11:Z) := 0. Definition P_id_0 := 0. Definition P_id_U12 (x11:Z) := 0. Definition P_id_a__U52 (x11:Z) (x12:Z) := 2* x12. Definition P_id_a__isNat (x11:Z) := 1. Definition P_id_U43 (x11:Z) := 2. Definition P_id_a__isNatIListKind (x11:Z) := 1. Definition P_id_a__U12 (x11:Z) := 1. Definition P_id_U31 (x11:Z) (x12:Z) := 1 + 3* x12. Definition P_id_a__length (x11:Z) := 1* x11. Definition P_id_a__U42 (x11:Z) (x12:Z) := 3. Definition P_id_U53 (x11:Z) := 1* x11. Definition P_id_and (x11:Z) (x12:Z) := 1* x11 + 1* x12. Definition P_id_cons (x11:Z) (x12:Z) := 1 + 3* x12. Definition P_id_U11 (x11:Z) (x12:Z) := 0. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 1 + 2* x13. Definition P_id_a__U22 (x11:Z) := 1* x11. Definition P_id_U42 (x11:Z) (x12:Z) := 1. Definition P_id_length (x11:Z) := 1* x11. Definition P_id_tt := 1. Definition P_id_U22 (x11:Z) := 1* x11. Definition P_id_s (x11:Z) := 1 + 1* x11. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 2 + 2* x11 + 2* x13. Definition P_id_U52 (x11:Z) (x12:Z) := 1* x12. Definition P_id_nil := 1. Definition P_id_zeros := 0. Definition P_id_isNatList (x11:Z) := 1* x11. Definition P_id_a__U53 (x11:Z) := 1* x11. Definition P_id_a__U31 (x11:Z) (x12:Z) := 1 + 3* x12. Definition P_id_isNatIList (x11:Z) := 2 + 2* x11. Definition P_id_a__isNatKind (x11:Z) := 1. Definition P_id_a__isNatList (x11:Z) := 1* x11. Definition P_id_U32 (x11:Z) := 0. Definition P_id_mark (x11:Z) := 1 + 2* x11. Definition P_id_a__U43 (x11:Z) := 3. Definition P_id_U61 (x11:Z) (x12:Z) := 1 + 1* x11 + 2* x12. Definition P_id_isNat (x11:Z) := 0. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12::x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12::x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12::x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12::x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12::x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12::x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12::x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12::x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic; assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 0. Definition P_id_A__U11 (x11:Z) (x12:Z) := 0. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 0. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__AND (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATLIST (x11:Z) := 0. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 1* x11. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 0. Definition P_id_A__U52 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILISTKIND (x11:Z) := 0. Definition P_id_A__U21 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 0. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61 . Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic; assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic; assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_large := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_large.wf . Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1 . Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_large) . clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1_large. Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1. Definition wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1 := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1.wf . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1. Qed. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_non_scc_2 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_non_scc_2_0 : forall x12 x5 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_non_scc_2 (algebra.Alg.Term algebra.F.id_mark (x5::nil)) (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_non_scc_2 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_non_scc_2 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large 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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_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. Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large . Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_non_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (fail)))))). Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_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_a__zeros := 2. Definition P_id_isNatKind (x11:Z) := 1* x11. Definition P_id_a__isNatIList (x11:Z) := 3 + 2* x11. Definition P_id_a__U21 (x11:Z) (x12:Z) := 0. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 1* x12 + 1* x13. Definition P_id_a__and (x11:Z) (x12:Z) := 2 + 2* x11 + 2* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 0. Definition P_id_U21 (x11:Z) (x12:Z) := 0. Definition P_id_a__U61 (x11:Z) (x12:Z) := 3 + 2* x12. Definition P_id_a__U32 (x11:Z) := 2. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_isNatIListKind (x11:Z) := 1 + 2* x11. Definition P_id_0 := 0. Definition P_id_U12 (x11:Z) := 0. Definition P_id_a__U52 (x11:Z) (x12:Z) := 0. Definition P_id_a__isNat (x11:Z) := 0. Definition P_id_U43 (x11:Z) := 0. Definition P_id_a__isNatIListKind (x11:Z) := 1 + 2* x11. Definition P_id_a__U12 (x11:Z) := 0. Definition P_id_U31 (x11:Z) (x12:Z) := 1* x12. Definition P_id_a__length (x11:Z) := 1 + 1* x11. Definition P_id_a__U42 (x11:Z) (x12:Z) := 0. Definition P_id_U53 (x11:Z) := 0. Definition P_id_and (x11:Z) (x12:Z) := 2 + 2* x11 + 2* x12. Definition P_id_cons (x11:Z) (x12:Z) := 2 + 2* x11 + 2* x12 . Definition P_id_U11 (x11:Z) (x12:Z) := 0. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_a__U22 (x11:Z) := 1* x11. Definition P_id_U42 (x11:Z) (x12:Z) := 0. Definition P_id_length (x11:Z) := 1 + 1* x11. Definition P_id_tt := 0. Definition P_id_U22 (x11:Z) := 1* x11. Definition P_id_s (x11:Z) := 1* x11. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 1* x12 + 1* x13. Definition P_id_U52 (x11:Z) (x12:Z) := 0. Definition P_id_nil := 0. Definition P_id_zeros := 0. Definition P_id_isNatList (x11:Z) := 0. Definition P_id_a__U53 (x11:Z) := 0. Definition P_id_a__U31 (x11:Z) (x12:Z) := 2 + 2* x12. Definition P_id_isNatIList (x11:Z) := 1 + 2* x11. Definition P_id_a__isNatKind (x11:Z) := 2* x11. Definition P_id_a__isNatList (x11:Z) := 0. Definition P_id_U32 (x11:Z) := 0. Definition P_id_mark (x11:Z) := 2 + 2* x11. Definition P_id_a__U43 (x11:Z) := 0. Definition P_id_U61 (x11:Z) (x12:Z) := 1 + 2* x12. Definition P_id_isNat (x11:Z) := 0. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13:: x12::x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12:: x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13:: x12::x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12:: x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12:: x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12::x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12:: x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12:: x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12:: x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12:: x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 0. Definition P_id_A__U11 (x11:Z) (x12:Z) := 0. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 0. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__AND (x11:Z) (x12:Z) := 2 + 1* x12. Definition P_id_A__ISNATLIST (x11:Z) := 0. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 2 + 1* x11. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 0. Definition P_id_A__U52 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILISTKIND (x11:Z) := 2* x11. Definition P_id_A__U21 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 0. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large.wf . Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large . Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large) . clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_large. Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_a__zeros := 2. Definition P_id_isNatKind (x11:Z) := 0. Definition P_id_a__isNatIList (x11:Z) := 2 + 2* x11. Definition P_id_a__U21 (x11:Z) (x12:Z) := 0. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 1 + 1* x13. Definition P_id_a__and (x11:Z) (x12:Z) := 2* x11 + 2* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 0. Definition P_id_U21 (x11:Z) (x12:Z) := 0. Definition P_id_a__U61 (x11:Z) (x12:Z) := 0. Definition P_id_a__U32 (x11:Z) := 2. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 2 + 2* x13. Definition P_id_isNatIListKind (x11:Z) := 0. Definition P_id_0 := 0. Definition P_id_U12 (x11:Z) := 0. Definition P_id_a__U52 (x11:Z) (x12:Z) := 2 + 2* x12. Definition P_id_a__isNat (x11:Z) := 0. Definition P_id_U43 (x11:Z) := 1. Definition P_id_a__isNatIListKind (x11:Z) := 0. Definition P_id_a__U12 (x11:Z) := 0. Definition P_id_U31 (x11:Z) (x12:Z) := 1 + 1* x12. Definition P_id_a__length (x11:Z) := 0. Definition P_id_a__U42 (x11:Z) (x12:Z) := 1 + 2* x12. Definition P_id_U53 (x11:Z) := 2 + 1* x11. Definition P_id_and (x11:Z) (x12:Z) := 2* x11 + 1* x12. Definition P_id_cons (x11:Z) (x12:Z) := 1 + 1* x12. Definition P_id_U11 (x11:Z) (x12:Z) := 0. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 2 + 2* x13 . Definition P_id_a__U22 (x11:Z) := 1* x11. Definition P_id_U42 (x11:Z) (x12:Z) := 1 + 1* x12. Definition P_id_length (x11:Z) := 0. Definition P_id_tt := 0. Definition P_id_U22 (x11:Z) := 1* x11. Definition P_id_s (x11:Z) := 2* x11. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 2 + 2* x13 . Definition P_id_U52 (x11:Z) (x12:Z) := 1 + 1* x12. Definition P_id_nil := 0. Definition P_id_zeros := 1. Definition P_id_isNatList (x11:Z) := 1* x11. Definition P_id_a__U53 (x11:Z) := 2 + 1* x11. Definition P_id_a__U31 (x11:Z) (x12:Z) := 2 + 1* x12. Definition P_id_isNatIList (x11:Z) := 2 + 1* x11. Definition P_id_a__isNatKind (x11:Z) := 0. Definition P_id_a__isNatList (x11:Z) := 2* x11. Definition P_id_U32 (x11:Z) := 1. Definition P_id_mark (x11:Z) := 2* x11. Definition P_id_a__U43 (x11:Z) := 1. Definition P_id_U61 (x11:Z) (x12:Z) := 0. Definition P_id_isNat (x11:Z) := 0. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13:: x12::x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12:: x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13:: x12::x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12:: x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12:: x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12:: x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12:: x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12:: x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12:: x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12:: x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 0. Definition P_id_A__U11 (x11:Z) (x12:Z) := 0. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 0. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__AND (x11:Z) (x12:Z) := 2* x12. Definition P_id_A__ISNATLIST (x11:Z) := 0. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 2* x11. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 0. Definition P_id_A__U52 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILISTKIND (x11:Z) := 0. Definition P_id_A__U21 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 0. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_strict lt. Proof. unfold Relation_Definitions.inclusion, lt in *. intros a b H;destruct H; match goal with | |- (Zwf.Zwf 0) _ (marked_measure (algebra.Alg.Term ?f ?l)) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_lt_compat_right (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[ idtac| apply marked_measure_star_monotonic; repeat ( apply algebra.EQT_ext.one_step_list_refl_trans_clos );(assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large le. Proof. unfold Relation_Definitions.inclusion, le, Zwf.Zwf in *. intros a b H;destruct H; match goal with | |- _ <= marked_measure (algebra.Alg.Term ?f ?l) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_trans (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[ idtac| apply marked_measure_star_monotonic; repeat ( apply algebra.EQT_ext.one_step_list_refl_trans_clos );(assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Definition wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large.wf . Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4 . Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large) . clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_strict_in_lt; econstructor eassumption)|| ((apply IHx;[econstructor eassumption| intros y' Hlt;apply IHx'; apply lt_le_compat with (1:=Hlt) ; apply DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4_large. Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4. Definition wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4 := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4.wf . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4) . intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_non_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_non_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' ))))). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4. Qed. Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large . Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_non_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_non_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_4; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (fail)))))))))). Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_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_a__zeros := 0. Definition P_id_isNatKind (x11:Z) := 1* x11. Definition P_id_a__isNatIList (x11:Z) := 2 + 2* x11. Definition P_id_a__U21 (x11:Z) (x12:Z) := 0. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 1 + 1* x12. Definition P_id_a__and (x11:Z) (x12:Z) := 1* x11 + 2* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 0. Definition P_id_U21 (x11:Z) (x12:Z) := 0. Definition P_id_a__U61 (x11:Z) (x12:Z) := 1 + 2* x12. Definition P_id_a__U32 (x11:Z) := 0. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_isNatIListKind (x11:Z) := 1* x11. Definition P_id_0 := 0. Definition P_id_U12 (x11:Z) := 1* x11. Definition P_id_a__U52 (x11:Z) (x12:Z) := 0. Definition P_id_a__isNat (x11:Z) := 0. Definition P_id_U43 (x11:Z) := 0. Definition P_id_a__isNatIListKind (x11:Z) := 1* x11. Definition P_id_a__U12 (x11:Z) := 1* x11. Definition P_id_U31 (x11:Z) (x12:Z) := 1* x11. Definition P_id_a__length (x11:Z) := 1 + 1* x11. Definition P_id_a__U42 (x11:Z) (x12:Z) := 0. Definition P_id_U53 (x11:Z) := 2* x11. Definition P_id_and (x11:Z) (x12:Z) := 1* x11 + 1* x12. Definition P_id_cons (x11:Z) (x12:Z) := 1* x11 + 2* x12. Definition P_id_U11 (x11:Z) (x12:Z) := 0. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_a__U22 (x11:Z) := 1* x11. Definition P_id_U42 (x11:Z) (x12:Z) := 0. Definition P_id_length (x11:Z) := 1 + 1* x11. Definition P_id_tt := 0. Definition P_id_U22 (x11:Z) := 1* x11. Definition P_id_s (x11:Z) := 1* x11. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 2 + 2* x12. Definition P_id_U52 (x11:Z) (x12:Z) := 0. Definition P_id_nil := 0. Definition P_id_zeros := 0. Definition P_id_isNatList (x11:Z) := 0. Definition P_id_a__U53 (x11:Z) := 2* x11. Definition P_id_a__U31 (x11:Z) (x12:Z) := 1* x11. Definition P_id_isNatIList (x11:Z) := 2 + 1* x11. Definition P_id_a__isNatKind (x11:Z) := 1* x11. Definition P_id_a__isNatList (x11:Z) := 0. Definition P_id_U32 (x11:Z) := 0. Definition P_id_mark (x11:Z) := 2* x11. Definition P_id_a__U43 (x11:Z) := 0. Definition P_id_U61 (x11:Z) (x12:Z) := 1 + 2* x12. Definition P_id_isNat (x11:Z) := 0. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13:: x12::x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12:: x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13:: x12::x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12:: x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12:: x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12:: x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12:: x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13:: x12::x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12:: x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13:: x12::x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12:: x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12:: x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 0. Definition P_id_A__U11 (x11:Z) (x12:Z) := 0. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 1* x11. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__AND (x11:Z) (x12:Z) := 1* x12. Definition P_id_A__ISNATLIST (x11:Z) := 0. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 1* x11. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 0. Definition P_id_A__U52 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILISTKIND (x11:Z) := 1* x11. Definition P_id_A__U21 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 0. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large.wf . Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large . Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large) . clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_large. Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_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_a__zeros := 0. Definition P_id_isNatKind (x11:Z) := 0. Definition P_id_a__isNatIList (x11:Z) := 0. Definition P_id_a__U21 (x11:Z) (x12:Z) := 2. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_a__and (x11:Z) (x12:Z) := 2* x11 + 2* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 1. Definition P_id_U21 (x11:Z) (x12:Z) := 2. Definition P_id_a__U61 (x11:Z) (x12:Z) := 0. Definition P_id_a__U32 (x11:Z) := 0. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_isNatIListKind (x11:Z) := 0. Definition P_id_0 := 0. Definition P_id_U12 (x11:Z) := 1 + 1* x11. Definition P_id_a__U52 (x11:Z) (x12:Z) := 0. Definition P_id_a__isNat (x11:Z) := 2. Definition P_id_U43 (x11:Z) := 0. Definition P_id_a__isNatIListKind (x11:Z) := 0. Definition P_id_a__U12 (x11:Z) := 1 + 1* x11. Definition P_id_U31 (x11:Z) (x12:Z) := 0. Definition P_id_a__length (x11:Z) := 0. Definition P_id_a__U42 (x11:Z) (x12:Z) := 0. Definition P_id_U53 (x11:Z) := 2* x11. Definition P_id_and (x11:Z) (x12:Z) := 2* x11 + 2* x12. Definition P_id_cons (x11:Z) (x12:Z) := 0. Definition P_id_U11 (x11:Z) (x12:Z) := 1. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_a__U22 (x11:Z) := 1* x11. Definition P_id_U42 (x11:Z) (x12:Z) := 0. Definition P_id_length (x11:Z) := 0. Definition P_id_tt := 0. Definition P_id_U22 (x11:Z) := 1* x11. Definition P_id_s (x11:Z) := 1* x11. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_U52 (x11:Z) (x12:Z) := 0. Definition P_id_nil := 1. Definition P_id_zeros := 0. Definition P_id_isNatList (x11:Z) := 0. Definition P_id_a__U53 (x11:Z) := 2* x11. Definition P_id_a__U31 (x11:Z) (x12:Z) := 0. Definition P_id_isNatIList (x11:Z) := 0. Definition P_id_a__isNatKind (x11:Z) := 0. Definition P_id_a__isNatList (x11:Z) := 0. Definition P_id_U32 (x11:Z) := 0. Definition P_id_mark (x11:Z) := 1* x11. Definition P_id_a__U43 (x11:Z) := 0. Definition P_id_U61 (x11:Z) (x12:Z) := 0. Definition P_id_isNat (x11:Z) := 2. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13:: x12::x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12:: x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13:: x12::x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12:: x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12:: x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12:: x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12:: x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13:: x12::x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12:: x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13:: x12::x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12:: x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12:: x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 0. Definition P_id_A__U11 (x11:Z) (x12:Z) := 0. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 0. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__AND (x11:Z) (x12:Z) := 2* x12. Definition P_id_A__ISNATLIST (x11:Z) := 0. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 2* x11. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 0. Definition P_id_A__U52 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILISTKIND (x11:Z) := 0. Definition P_id_A__U21 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 0. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large.wf . Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 . Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large) . clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1_large. Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1. Definition wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1.wf . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1. Qed. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_non_scc_2 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_non_scc_2_0 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_non_scc_2 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_non_scc_2 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_non_scc_2 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large 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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_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_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_0 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3 (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_1 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3 (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large_0 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) . Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_strict_0 : forall x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_s (x1::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_strict (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large_non_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large_non_scc_1_0 : forall x12 x1 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large_non_scc_1 (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large_non_scc_1 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large_non_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large . Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (fail))). Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_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_a__zeros := 1. Definition P_id_isNatKind (x11:Z) := 0. Definition P_id_a__isNatIList (x11:Z) := 1 + 3* x11. Definition P_id_a__U21 (x11:Z) (x12:Z) := 1. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 1* x11. Definition P_id_a__and (x11:Z) (x12:Z) := 1* x11 + 1* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 1. Definition P_id_U21 (x11:Z) (x12:Z) := 1. Definition P_id_a__U61 (x11:Z) (x12:Z) := 1 + 1* x11 + 2* x12. Definition P_id_a__U32 (x11:Z) := 1* x11. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 3* x13. Definition P_id_isNatIListKind (x11:Z) := 0. Definition P_id_0 := 0. Definition P_id_U12 (x11:Z) := 0. Definition P_id_a__U52 (x11:Z) (x12:Z) := 2* x12. Definition P_id_a__isNat (x11:Z) := 1. Definition P_id_U43 (x11:Z) := 0. Definition P_id_a__isNatIListKind (x11:Z) := 1. Definition P_id_a__U12 (x11:Z) := 1. Definition P_id_U31 (x11:Z) (x12:Z) := 1* x12. Definition P_id_a__length (x11:Z) := 1* x11. Definition P_id_a__U42 (x11:Z) (x12:Z) := 1. Definition P_id_U53 (x11:Z) := 1* x11. Definition P_id_and (x11:Z) (x12:Z) := 1* x11 + 1* x12. Definition P_id_cons (x11:Z) (x12:Z) := 1 + 3* x12. Definition P_id_U11 (x11:Z) (x12:Z) := 1. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 3* x13. Definition P_id_a__U22 (x11:Z) := 1. Definition P_id_U42 (x11:Z) (x12:Z) := 0. Definition P_id_length (x11:Z) := 1* x11. Definition P_id_tt := 1. Definition P_id_U22 (x11:Z) := 0. Definition P_id_s (x11:Z) := 1 + 1* x11. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 1* x11. Definition P_id_U52 (x11:Z) (x12:Z) := 2* x12. Definition P_id_nil := 1. Definition P_id_zeros := 0. Definition P_id_isNatList (x11:Z) := 1* x11. Definition P_id_a__U53 (x11:Z) := 1* x11. Definition P_id_a__U31 (x11:Z) (x12:Z) := 1* x12. Definition P_id_isNatIList (x11:Z) := 3* x11. Definition P_id_a__isNatKind (x11:Z) := 1. Definition P_id_a__isNatList (x11:Z) := 1* x11. Definition P_id_U32 (x11:Z) := 1* x11. Definition P_id_mark (x11:Z) := 1 + 1* x11. Definition P_id_a__U43 (x11:Z) := 1. Definition P_id_U61 (x11:Z) (x12:Z) := 1 + 1* x11 + 2* x12. Definition P_id_isNat (x11:Z) := 0. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13:: x12::x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12:: x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13:: x12::x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12:: x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12:: x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12:: x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12:: x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13:: x12::x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12:: x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13:: x12::x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12:: x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12:: x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 0. Definition P_id_A__U11 (x11:Z) (x12:Z) := 0. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 0. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__AND (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATLIST (x11:Z) := 0. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 0. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 1* x11. Definition P_id_A__U52 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILISTKIND (x11:Z) := 0. Definition P_id_A__U21 (x11:Z) (x12:Z) := 1* x12. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 0. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large.wf . Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3 . Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large) . clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3_large. Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3. Definition wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3 := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3.wf . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_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)|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_non_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' )))). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3. Qed. Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large . Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_non_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (fail)))))))). Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_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_a__zeros := 0. Definition P_id_isNatKind (x11:Z) := 0. Definition P_id_a__isNatIList (x11:Z) := 0. Definition P_id_a__U21 (x11:Z) (x12:Z) := 1 + 2* x11. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_a__and (x11:Z) (x12:Z) := 2* x11 + 2* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 0. Definition P_id_U21 (x11:Z) (x12:Z) := 1 + 2* x11. Definition P_id_a__U61 (x11:Z) (x12:Z) := 0. Definition P_id_a__U32 (x11:Z) := 0. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_isNatIListKind (x11:Z) := 0. Definition P_id_0 := 0. Definition P_id_U12 (x11:Z) := 1* x11. Definition P_id_a__U52 (x11:Z) (x12:Z) := 0. Definition P_id_a__isNat (x11:Z) := 1. Definition P_id_U43 (x11:Z) := 0. Definition P_id_a__isNatIListKind (x11:Z) := 0. Definition P_id_a__U12 (x11:Z) := 1* x11. Definition P_id_U31 (x11:Z) (x12:Z) := 0. Definition P_id_a__length (x11:Z) := 0. Definition P_id_a__U42 (x11:Z) (x12:Z) := 0. Definition P_id_U53 (x11:Z) := 2* x11. Definition P_id_and (x11:Z) (x12:Z) := 2* x11 + 2* x12. Definition P_id_cons (x11:Z) (x12:Z) := 0. Definition P_id_U11 (x11:Z) (x12:Z) := 0. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_a__U22 (x11:Z) := 1* x11. Definition P_id_U42 (x11:Z) (x12:Z) := 0. Definition P_id_length (x11:Z) := 0. Definition P_id_tt := 0. Definition P_id_U22 (x11:Z) := 1* x11. Definition P_id_s (x11:Z) := 2* x11. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_U52 (x11:Z) (x12:Z) := 0. Definition P_id_nil := 0. Definition P_id_zeros := 0. Definition P_id_isNatList (x11:Z) := 0. Definition P_id_a__U53 (x11:Z) := 2* x11. Definition P_id_a__U31 (x11:Z) (x12:Z) := 0. Definition P_id_isNatIList (x11:Z) := 0. Definition P_id_a__isNatKind (x11:Z) := 0. Definition P_id_a__isNatList (x11:Z) := 0. Definition P_id_U32 (x11:Z) := 0. Definition P_id_mark (x11:Z) := 1* x11. Definition P_id_a__U43 (x11:Z) := 0. Definition P_id_U61 (x11:Z) (x12:Z) := 0. Definition P_id_isNat (x11:Z) := 1. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13::x12:: x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12:: x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13::x12:: x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12:: x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12:: x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12:: x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12:: x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13:: x12::x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12:: x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13:: x12::x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12:: x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12:: x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 0. Definition P_id_A__U11 (x11:Z) (x12:Z) := 0. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 0. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__AND (x11:Z) (x12:Z) := 1* x12. Definition P_id_A__ISNATLIST (x11:Z) := 0. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 1* x11. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 0. Definition P_id_A__U52 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILISTKIND (x11:Z) := 0. Definition P_id_A__U21 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 0. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large.wf . Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 . Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large) . clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_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_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7_large. Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7. Definition wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7.wf . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7) . intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_6; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_5; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_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' ))))))). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7. Qed. Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large . Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_6; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_5; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_4; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_7; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_6; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_5; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_4; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (fail)))))))))))))))). Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_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_a__zeros := 0. Definition P_id_isNatKind (x11:Z) := 0. Definition P_id_a__isNatIList (x11:Z) := 0. Definition P_id_a__U21 (x11:Z) (x12:Z) := 1* x11 + 1* x12. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_a__and (x11:Z) (x12:Z) := 1* x11 + 2* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 1 + 2* x12. Definition P_id_U21 (x11:Z) (x12:Z) := 1* x11 + 1* x12. Definition P_id_a__U61 (x11:Z) (x12:Z) := 1 + 2* x12. Definition P_id_a__U32 (x11:Z) := 0. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 1 + 1* x11 + 2* x12 + 2* x13. Definition P_id_isNatIListKind (x11:Z) := 0. Definition P_id_0 := 0. Definition P_id_U12 (x11:Z) := 1* x11. Definition P_id_a__U52 (x11:Z) (x12:Z) := 1 + 2* x11 + 2* x12. Definition P_id_a__isNat (x11:Z) := 1* x11. Definition P_id_U43 (x11:Z) := 0. Definition P_id_a__isNatIListKind (x11:Z) := 0. Definition P_id_a__U12 (x11:Z) := 1* x11. Definition P_id_U31 (x11:Z) (x12:Z) := 1* x11. Definition P_id_a__length (x11:Z) := 1 + 2* x11. Definition P_id_a__U42 (x11:Z) (x12:Z) := 0. Definition P_id_U53 (x11:Z) := 1* x11. Definition P_id_and (x11:Z) (x12:Z) := 1* x11 + 2* x12. Definition P_id_cons (x11:Z) (x12:Z) := 1* x11 + 1* x12. Definition P_id_U11 (x11:Z) (x12:Z) := 1 + 2* x12. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 1 + 1* x11 + 2* x12 + 2* x13. Definition P_id_a__U22 (x11:Z) := 1* x11. Definition P_id_U42 (x11:Z) (x12:Z) := 0. Definition P_id_length (x11:Z) := 1 + 2* x11. Definition P_id_tt := 0. Definition P_id_U22 (x11:Z) := 1* x11. Definition P_id_s (x11:Z) := 1* x11. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_U52 (x11:Z) (x12:Z) := 1 + 2* x11 + 2* x12. Definition P_id_nil := 1. Definition P_id_zeros := 0. Definition P_id_isNatList (x11:Z) := 1 + 2* x11. Definition P_id_a__U53 (x11:Z) := 1* x11. Definition P_id_a__U31 (x11:Z) (x12:Z) := 1* x11. Definition P_id_isNatIList (x11:Z) := 0. Definition P_id_a__isNatKind (x11:Z) := 0. Definition P_id_a__isNatList (x11:Z) := 1 + 2* x11. Definition P_id_U32 (x11:Z) := 0. Definition P_id_mark (x11:Z) := 1* x11. Definition P_id_a__U43 (x11:Z) := 0. Definition P_id_U61 (x11:Z) (x12:Z) := 1 + 2* x12. Definition P_id_isNat (x11:Z) := 1* x11. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13::x12:: x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12:: x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13::x12:: x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12:: x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12:: x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12:: x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12:: x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12:: x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12:: x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12:: x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 0. Definition P_id_A__U11 (x11:Z) (x12:Z) := 1 + 2* x12. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 0. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 1 + 1* x12 + 2* x13. Definition P_id_A__AND (x11:Z) (x12:Z) := 1* x12. Definition P_id_A__ISNATLIST (x11:Z) := 1 + 2* x11. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 1* x11. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 1* x11. Definition P_id_A__U52 (x11:Z) (x12:Z) := 1 + 2* x12. Definition P_id_A__ISNATILISTKIND (x11:Z) := 0. Definition P_id_A__U21 (x11:Z) (x12:Z) := 1* x12. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 0. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_2_large_large_scc_1_large_large_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large.wf . Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large) . clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large_large. Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_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_a__zeros := 3. Definition P_id_isNatKind (x11:Z) := 0. Definition P_id_a__isNatIList (x11:Z) := 1* x11. Definition P_id_a__U21 (x11:Z) (x12:Z) := 2* x11. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 1 + 1* x13. Definition P_id_a__and (x11:Z) (x12:Z) := 2* x11 + 2* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 0. Definition P_id_U21 (x11:Z) (x12:Z) := 2* x11. Definition P_id_a__U61 (x11:Z) (x12:Z) := 0. Definition P_id_a__U32 (x11:Z) := 0. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 2* x11. Definition P_id_isNatIListKind (x11:Z) := 0. Definition P_id_0 := 0. Definition P_id_U12 (x11:Z) := 2* x11. Definition P_id_a__U52 (x11:Z) (x12:Z) := 2* x11. Definition P_id_a__isNat (x11:Z) := 0. Definition P_id_U43 (x11:Z) := 1 + 1* x11. Definition P_id_a__isNatIListKind (x11:Z) := 0. Definition P_id_a__U12 (x11:Z) := 2* x11. Definition P_id_U31 (x11:Z) (x12:Z) := 2* x11 + 1* x12. Definition P_id_a__length (x11:Z) := 0. Definition P_id_a__U42 (x11:Z) (x12:Z) := 1 + 1* x12. Definition P_id_U53 (x11:Z) := 1* x11. Definition P_id_and (x11:Z) (x12:Z) := 2* x11 + 1* x12. Definition P_id_cons (x11:Z) (x12:Z) := 1 + 1* x12. Definition P_id_U11 (x11:Z) (x12:Z) := 0. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 2* x11. Definition P_id_a__U22 (x11:Z) := 2* x11. Definition P_id_U42 (x11:Z) (x12:Z) := 1 + 1* x12. Definition P_id_length (x11:Z) := 0. Definition P_id_tt := 0. Definition P_id_U22 (x11:Z) := 2* x11. Definition P_id_s (x11:Z) := 2* x11. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 1 + 1* x13. Definition P_id_U52 (x11:Z) (x12:Z) := 2* x11. Definition P_id_nil := 2. Definition P_id_zeros := 2. Definition P_id_isNatList (x11:Z) := 0. Definition P_id_a__U53 (x11:Z) := 1* x11. Definition P_id_a__U31 (x11:Z) (x12:Z) := 2* x11 + 1* x12. Definition P_id_isNatIList (x11:Z) := 1* x11. Definition P_id_a__isNatKind (x11:Z) := 0. Definition P_id_a__isNatList (x11:Z) := 0. Definition P_id_U32 (x11:Z) := 0. Definition P_id_mark (x11:Z) := 2* x11. Definition P_id_a__U43 (x11:Z) := 1 + 1* x11. Definition P_id_U61 (x11:Z) (x12:Z) := 0. Definition P_id_isNat (x11:Z) := 0. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) -> (0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13::x12:: x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12:: x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13::x12:: x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12:: x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12:: x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12:: x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12:: x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12:: x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12:: x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12:: x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 0. Definition P_id_A__U11 (x11:Z) (x12:Z) := 0. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 0. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__AND (x11:Z) (x12:Z) := 2* x12. Definition P_id_A__ISNATLIST (x11:Z) := 0. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 2* x11. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 0. Definition P_id_A__U52 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILISTKIND (x11:Z) := 0. Definition P_id_A__U21 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 0. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12::x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_2_large_large_scc_1_large_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large.wf. Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large . Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large) . clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large_large. Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_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_a__zeros := 0. Definition P_id_isNatKind (x11:Z) := 0. Definition P_id_a__isNatIList (x11:Z) := 0. Definition P_id_a__U21 (x11:Z) (x12:Z) := 2* x11 + 1* x12. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 3* x11. Definition P_id_a__and (x11:Z) (x12:Z) := 1* x11 + 2* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 1 + 2* x11 + 1* x12. Definition P_id_U21 (x11:Z) (x12:Z) := 2* x11 + 1* x12. Definition P_id_a__U61 (x11:Z) (x12:Z) := 1 + 2* x12. Definition P_id_a__U32 (x11:Z) := 0. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 1* x11 + 1* x12 + 1* x13. Definition P_id_isNatIListKind (x11:Z) := 0. Definition P_id_0 := 0. Definition P_id_U12 (x11:Z) := 1* x11. Definition P_id_a__U52 (x11:Z) (x12:Z) := 2* x11 + 2* x12. Definition P_id_a__isNat (x11:Z) := 1* x11. Definition P_id_U43 (x11:Z) := 1* x11. Definition P_id_a__isNatIListKind (x11:Z) := 0. Definition P_id_a__U12 (x11:Z) := 1* x11. Definition P_id_U31 (x11:Z) (x12:Z) := 0. Definition P_id_a__length (x11:Z) := 1 + 1* x11. Definition P_id_a__U42 (x11:Z) (x12:Z) := 0. Definition P_id_U53 (x11:Z) := 1* x11. Definition P_id_and (x11:Z) (x12:Z) := 1* x11 + 2* x12. Definition P_id_cons (x11:Z) (x12:Z) := 2* x11 + 2* x12. Definition P_id_U11 (x11:Z) (x12:Z) := 1 + 2* x11 + 1* x12. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 1* x11 + 2* x12 + 2* x13. Definition P_id_a__U22 (x11:Z) := 1* x11. Definition P_id_U42 (x11:Z) (x12:Z) := 0. Definition P_id_length (x11:Z) := 1 + 1* x11. Definition P_id_tt := 0. Definition P_id_U22 (x11:Z) := 1* x11. Definition P_id_s (x11:Z) := 1* x11. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 3* x11. Definition P_id_U52 (x11:Z) (x12:Z) := 2* x11 + 2* x12. Definition P_id_nil := 1. Definition P_id_zeros := 0. Definition P_id_isNatList (x11:Z) := 1* x11. Definition P_id_a__U53 (x11:Z) := 1* x11. Definition P_id_a__U31 (x11:Z) (x12:Z) := 0. Definition P_id_isNatIList (x11:Z) := 0. Definition P_id_a__isNatKind (x11:Z) := 0. Definition P_id_a__isNatList (x11:Z) := 1* x11. Definition P_id_U32 (x11:Z) := 0. Definition P_id_mark (x11:Z) := 2* x11. Definition P_id_a__U43 (x11:Z) := 1* x11. Definition P_id_U61 (x11:Z) (x12:Z) := 1 + 2* x12. Definition P_id_isNat (x11:Z) := 1* x11. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13::x12:: x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12::x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13::x12:: x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12::x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12::x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12:: x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12::x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12::x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12::x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12::x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 0. Definition P_id_A__U11 (x11:Z) (x12:Z) := 2. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 2. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 2. Definition P_id_A__AND (x11:Z) (x12:Z) := 2 + 2* x12. Definition P_id_A__ISNATLIST (x11:Z) := 2. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 2 + 2* x11. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 2. Definition P_id_A__U52 (x11:Z) (x12:Z) := 2. Definition P_id_A__ISNATILISTKIND (x11:Z) := 2. Definition P_id_A__U21 (x11:Z) (x12:Z) := 2. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 0. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13:: x12::x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13:: x12::x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_2_large_large_scc_1_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large.wf. Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 . Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_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_12_large_scc_2_large_large_scc_1_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1_large. Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1. Definition wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1.wf. Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large.DP_R_xml_0_scc_12_large_scc_2_large_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_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_12_large_scc_2_large_large_scc_1. Qed. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_2 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_2_0 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_2 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_2 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_2 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large.DP_R_xml_0_scc_12_large_scc_2_large_large 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_12_large_scc_2_large_large_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_scc_12_large_scc_2_large_large_non_scc_3 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_3_0 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_3 (algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_3 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_3 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large.DP_R_xml_0_scc_12_large_scc_2_large_large 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_12_large_scc_2_large_large_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_scc_12_large_scc_2_large_large_non_scc_4 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_4_0 : forall x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_4 (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x2::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_4 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_4 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large.DP_R_xml_0_scc_12_large_scc_2_large_large 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_12_large_scc_2_large_large_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_scc_12_large_scc_2_large_large_non_scc_5 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_5_0 : forall x12 x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_5 (algebra.Alg.Term algebra.F.id_a__isNatList (x2::nil)) (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_5 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_5 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large.DP_R_xml_0_scc_12_large_scc_2_large_large 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_12_large_scc_2_large_large_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_scc_12_large_scc_2_large_large_non_scc_6 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_6_0 : forall x2 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x2 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_6 (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_a__isNatIListKind (x2::nil)):: x2::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_6 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_6 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large.DP_R_xml_0_scc_12_large_scc_2_large_large 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_12_large_scc_2_large_large_non_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_12_large_scc_2_large_large_non_scc_7 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_7_0 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_7 (algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_7 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_7 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large.DP_R_xml_0_scc_12_large_scc_2_large_large 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_12_large_scc_2_large_large_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_scc_12_large_scc_2_large_large_scc_8 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_0 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil)):: x3::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_1 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8 (algebra.Alg.Term algebra.F.id_a__isNatIList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_2 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8 (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: (algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)):: x1::x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_0 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_1 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large (algebra.Alg.Term algebra.F.id_a__U41 (( algebra.Alg.Term algebra.F.id_a__and (( algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil)):: ( algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil)):: x1:: x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) . Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_strict_0 : forall x12 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_strict (algebra.Alg.Term algebra.F.id_a__isNatIList (x3::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) . Module WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_non_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_non_scc_1_0 : forall x12 x1 x13 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tt nil) x13) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x12) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_non_scc_1 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isNat (x1::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12::x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_non_scc_1 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_non_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_non_scc_2 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_non_scc_2_0 : forall x1 x3 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x1::x3::nil)) x11) -> DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_non_scc_2 (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isNatKind (x1::nil))::(algebra.Alg.Term algebra.F.id_isNatIListKind (x3::nil))::nil))::x1::x3::nil)) (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) . Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_non_scc_2 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_non_scc_2 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large 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_12_large_scc_2_large_large_scc_8_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large . Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_non_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||(fail)))). Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_a__zeros := 1. Definition P_id_isNatKind (x11:Z) := 1. Definition P_id_a__isNatIList (x11:Z) := 2. Definition P_id_a__U21 (x11:Z) (x12:Z) := 1. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 1. Definition P_id_a__and (x11:Z) (x12:Z) := 1 + 2* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 1. Definition P_id_U21 (x11:Z) (x12:Z) := 0. Definition P_id_a__U61 (x11:Z) (x12:Z) := 1. Definition P_id_a__U32 (x11:Z) := 2. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_isNatIListKind (x11:Z) := 0. Definition P_id_0 := 1. Definition P_id_U12 (x11:Z) := 1. Definition P_id_a__U52 (x11:Z) (x12:Z) := 1. Definition P_id_a__isNat (x11:Z) := 1* x11. Definition P_id_U43 (x11:Z) := 0. Definition P_id_a__isNatIListKind (x11:Z) := 1. Definition P_id_a__U12 (x11:Z) := 1. Definition P_id_U31 (x11:Z) (x12:Z) := 2. Definition P_id_a__length (x11:Z) := 1. Definition P_id_a__U42 (x11:Z) (x12:Z) := 1. Definition P_id_U53 (x11:Z) := 1. Definition P_id_and (x11:Z) (x12:Z) := 2* x12. Definition P_id_cons (x11:Z) (x12:Z) := 1* x11 + 1* x12. Definition P_id_U11 (x11:Z) (x12:Z) := 1. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 1. Definition P_id_a__U22 (x11:Z) := 1. Definition P_id_U42 (x11:Z) (x12:Z) := 0. Definition P_id_length (x11:Z) := 1. Definition P_id_tt := 1. Definition P_id_U22 (x11:Z) := 1. Definition P_id_s (x11:Z) := 1. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 1. Definition P_id_U52 (x11:Z) (x12:Z) := 0. Definition P_id_nil := 0. Definition P_id_zeros := 0. Definition P_id_isNatList (x11:Z) := 0. Definition P_id_a__U53 (x11:Z) := 1. Definition P_id_a__U31 (x11:Z) (x12:Z) := 2. Definition P_id_isNatIList (x11:Z) := 2. Definition P_id_a__isNatKind (x11:Z) := 2. Definition P_id_a__isNatList (x11:Z) := 1. Definition P_id_U32 (x11:Z) := 2. Definition P_id_mark (x11:Z) := 1 + 1* x11. Definition P_id_a__U43 (x11:Z) := 1. Definition P_id_U61 (x11:Z) (x12:Z) := 0. Definition P_id_isNat (x11:Z) := 1* x11. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13::x12:: x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12::x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13::x12:: x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12::x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12::x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12:: x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12::x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12::x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12::x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12::x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 0. Definition P_id_A__U11 (x11:Z) (x12:Z) := 0. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 1* x12 + 2* x13. Definition P_id_A__ISNATKIND (x11:Z) := 0. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__AND (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATLIST (x11:Z) := 0. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 0. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 1* x11 + 2* x12. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 0. Definition P_id_A__U52 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILISTKIND (x11:Z) := 0. Definition P_id_A__U21 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILIST (x11:Z) := 2* x11. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 0. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13:: x12::x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13:: x12::x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_2_large_large_scc_8_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_strict lt. Proof. unfold Relation_Definitions.inclusion, lt in *. intros a b H;destruct H; match goal with | |- (Zwf.Zwf 0) _ (marked_measure (algebra.Alg.Term ?f ?l)) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_lt_compat_right (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos );(assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large le. Proof. unfold Relation_Definitions.inclusion, le, Zwf.Zwf in *. intros a b H;destruct H; match goal with | |- _ <= marked_measure (algebra.Alg.Term ?f ?l) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_trans (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos );(assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large.wf. Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large_large.DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8 . Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_strict_in_lt; econstructor eassumption)|| ((apply IHx;[econstructor eassumption| intros y' Hlt;apply IHx';apply lt_le_compat with (1:=Hlt) ; apply DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8_large. Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8. Definition wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8 := WF_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8.wf. Lemma acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8 : forall x y, (DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8 x y) -> Acc WF_DP_R_xml_0_scc_12_large_scc_2_large.DP_R_xml_0_scc_12_large_scc_2_large_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8). intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_7; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_6; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_4; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_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' )))))))). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8. Qed. Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2_large.DP_R_xml_0_scc_12_large_scc_2_large_large . Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_7; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_6; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_5; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_4; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_8; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_7; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_6; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_5; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_4; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_2_large_large_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (fail)))))))))))))))))). Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_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_a__zeros := 0. Definition P_id_isNatKind (x11:Z) := 0. Definition P_id_a__isNatIList (x11:Z) := 1. Definition P_id_a__U21 (x11:Z) (x12:Z) := 1* x11. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 1 + 1* x11. Definition P_id_a__and (x11:Z) (x12:Z) := 2* x11 + 2* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 1* x11. Definition P_id_U21 (x11:Z) (x12:Z) := 1* x11. Definition P_id_a__U61 (x11:Z) (x12:Z) := 0. Definition P_id_a__U32 (x11:Z) := 1 + 2* x11. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 2* x11. Definition P_id_isNatIListKind (x11:Z) := 0. Definition P_id_0 := 0. Definition P_id_U12 (x11:Z) := 2* x11. Definition P_id_a__U52 (x11:Z) (x12:Z) := 2* x11. Definition P_id_a__isNat (x11:Z) := 0. Definition P_id_U43 (x11:Z) := 1* x11. Definition P_id_a__isNatIListKind (x11:Z) := 0. Definition P_id_a__U12 (x11:Z) := 2* x11. Definition P_id_U31 (x11:Z) (x12:Z) := 1 + 2* x11. Definition P_id_a__length (x11:Z) := 0. Definition P_id_a__U42 (x11:Z) (x12:Z) := 1 + 1* x11. Definition P_id_U53 (x11:Z) := 2* x11. Definition P_id_and (x11:Z) (x12:Z) := 2* x11 + 2* x12. Definition P_id_cons (x11:Z) (x12:Z) := 0. Definition P_id_U11 (x11:Z) (x12:Z) := 1* x11. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 2* x11. Definition P_id_a__U22 (x11:Z) := 1* x11. Definition P_id_U42 (x11:Z) (x12:Z) := 1 + 1* x11. Definition P_id_length (x11:Z) := 0. Definition P_id_tt := 0. Definition P_id_U22 (x11:Z) := 1* x11. Definition P_id_s (x11:Z) := 1* x11. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 1 + 1* x11. Definition P_id_U52 (x11:Z) (x12:Z) := 2* x11. Definition P_id_nil := 1. Definition P_id_zeros := 0. Definition P_id_isNatList (x11:Z) := 0. Definition P_id_a__U53 (x11:Z) := 2* x11. Definition P_id_a__U31 (x11:Z) (x12:Z) := 1 + 2* x11. Definition P_id_isNatIList (x11:Z) := 1. Definition P_id_a__isNatKind (x11:Z) := 0. Definition P_id_a__isNatList (x11:Z) := 0. Definition P_id_U32 (x11:Z) := 1 + 2* x11. Definition P_id_mark (x11:Z) := 2* x11. Definition P_id_a__U43 (x11:Z) := 1* x11. Definition P_id_U61 (x11:Z) (x12:Z) := 0. Definition P_id_isNat (x11:Z) := 0. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13::x12:: x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12::x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13::x12:: x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12::x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12::x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12::x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12::x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12::x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12::x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12::x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 0. Definition P_id_A__U11 (x11:Z) (x12:Z) := 0. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 0. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__AND (x11:Z) (x12:Z) := 2* x12. Definition P_id_A__ISNATLIST (x11:Z) := 0. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 2* x11. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 0. Definition P_id_A__U52 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILISTKIND (x11:Z) := 0. Definition P_id_A__U21 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 0. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13:: x12::x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13:: x12::x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_2_large_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_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_12_large_scc_2_large_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_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_12_large_scc_2_large_large := WF_DP_R_xml_0_scc_12_large_scc_2_large_large.wf. Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large_scc_2.DP_R_xml_0_scc_12_large_scc_2_large . Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx';apply DP_R_xml_0_scc_12_large_scc_2_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_12_large_scc_2_large_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_12_large_scc_2_large_large. Qed. End WF_DP_R_xml_0_scc_12_large_scc_2_large. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_a__zeros := 3. Definition P_id_isNatKind (x11:Z) := 0. Definition P_id_a__isNatIList (x11:Z) := 0. Definition P_id_a__U21 (x11:Z) (x12:Z) := 2* x11. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 1* x11. Definition P_id_a__and (x11:Z) (x12:Z) := 2* x11 + 2* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 2* x11. Definition P_id_U21 (x11:Z) (x12:Z) := 2* x11. Definition P_id_a__U61 (x11:Z) (x12:Z) := 0. Definition P_id_a__U32 (x11:Z) := 2* x11. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 2* x11. Definition P_id_isNatIListKind (x11:Z) := 0. Definition P_id_0 := 0. Definition P_id_U12 (x11:Z) := 1* x11. Definition P_id_a__U52 (x11:Z) (x12:Z) := 2* x11. Definition P_id_a__isNat (x11:Z) := 0. Definition P_id_U43 (x11:Z) := 1* x11. Definition P_id_a__isNatIListKind (x11:Z) := 0. Definition P_id_a__U12 (x11:Z) := 1* x11. Definition P_id_U31 (x11:Z) (x12:Z) := 2* x11. Definition P_id_a__length (x11:Z) := 0. Definition P_id_a__U42 (x11:Z) (x12:Z) := 2* x11. Definition P_id_U53 (x11:Z) := 1* x11. Definition P_id_and (x11:Z) (x12:Z) := 2* x11 + 2* x12. Definition P_id_cons (x11:Z) (x12:Z) := 3 + 1* x11. Definition P_id_U11 (x11:Z) (x12:Z) := 2* x11. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 2* x11. Definition P_id_a__U22 (x11:Z) := 1* x11. Definition P_id_U42 (x11:Z) (x12:Z) := 2* x11. Definition P_id_length (x11:Z) := 0. Definition P_id_tt := 0. Definition P_id_U22 (x11:Z) := 1* x11. Definition P_id_s (x11:Z) := 1* x11. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 1* x11. Definition P_id_U52 (x11:Z) (x12:Z) := 2* x11. Definition P_id_nil := 0. Definition P_id_zeros := 3. Definition P_id_isNatList (x11:Z) := 0. Definition P_id_a__U53 (x11:Z) := 1* x11. Definition P_id_a__U31 (x11:Z) (x12:Z) := 2* x11. Definition P_id_isNatIList (x11:Z) := 0. Definition P_id_a__isNatKind (x11:Z) := 0. Definition P_id_a__isNatList (x11:Z) := 0. Definition P_id_U32 (x11:Z) := 2* x11. Definition P_id_mark (x11:Z) := 1* x11. Definition P_id_a__U43 (x11:Z) := 1* x11. Definition P_id_U61 (x11:Z) (x12:Z) := 0. Definition P_id_isNat (x11:Z) := 0. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13::x12:: x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12::x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13::x12:: x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12::x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12::x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12::x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12::x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12::x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12::x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12::x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 0. Definition P_id_A__U11 (x11:Z) (x12:Z) := 0. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 0. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__AND (x11:Z) (x12:Z) := 1* x12. Definition P_id_A__ISNATLIST (x11:Z) := 0. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 1* x11. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 0. Definition P_id_A__U52 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILISTKIND (x11:Z) := 0. Definition P_id_A__U21 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 0. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13:: x12::x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13:: x12::x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_large_scc_2_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_strict lt. Proof. unfold Relation_Definitions.inclusion, lt in *. intros a b H;destruct H; match goal with | |- (Zwf.Zwf 0) _ (marked_measure (algebra.Alg.Term ?f ?l)) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_lt_compat_right (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma DP_R_xml_0_scc_12_large_scc_2_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_large_scc_2_large le. Proof. unfold Relation_Definitions.inclusion, le, Zwf.Zwf in *. intros a b H;destruct H; match goal with | |- _ <= marked_measure (algebra.Alg.Term ?f ?l) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_trans (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition wf_DP_R_xml_0_scc_12_large_scc_2_large := WF_DP_R_xml_0_scc_12_large_scc_2_large.wf. Lemma wf : well_founded WF_DP_R_xml_0_scc_12_large.DP_R_xml_0_scc_12_large_scc_2. Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx';apply DP_R_xml_0_scc_12_large_scc_2_strict_in_lt; econstructor eassumption)|| ((apply IHx;[econstructor eassumption| intros y' Hlt;apply IHx';apply lt_le_compat with (1:=Hlt) ; apply DP_R_xml_0_scc_12_large_scc_2_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_12_large_scc_2_large. Qed. End WF_DP_R_xml_0_scc_12_large_scc_2. Definition wf_DP_R_xml_0_scc_12_large_scc_2 := WF_DP_R_xml_0_scc_12_large_scc_2.wf. Lemma acc_DP_R_xml_0_scc_12_large_scc_2 : forall x y, (DP_R_xml_0_scc_12_large_scc_2 x y) -> Acc WF_DP_R_xml_0_scc_12.DP_R_xml_0_scc_12_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large_scc_2). intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((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_12_large_scc_2. Qed. Lemma wf : well_founded WF_DP_R_xml_0_scc_12.DP_R_xml_0_scc_12_large. Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_12_large_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12_large_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||(fail)))). Qed. End WF_DP_R_xml_0_scc_12_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_a__zeros := 0. Definition P_id_isNatKind (x11:Z) := 0. Definition P_id_a__isNatIList (x11:Z) := 0. Definition P_id_a__U21 (x11:Z) (x12:Z) := 2* x11. Definition P_id_U41 (x11:Z) (x12:Z) (x13:Z) := 2* x11. Definition P_id_a__and (x11:Z) (x12:Z) := 2* x11 + 1* x12. Definition P_id_a__U11 (x11:Z) (x12:Z) := 1* x11. Definition P_id_U21 (x11:Z) (x12:Z) := 2* x11. Definition P_id_a__U61 (x11:Z) (x12:Z) := 1 + 2* x11 + 1* x12. Definition P_id_a__U32 (x11:Z) := 2* x11. Definition P_id_U51 (x11:Z) (x12:Z) (x13:Z) := 1* x11. Definition P_id_isNatIListKind (x11:Z) := 0. Definition P_id_0 := 0. Definition P_id_U12 (x11:Z) := 1* x11. Definition P_id_a__U52 (x11:Z) (x12:Z) := 2* x11. Definition P_id_a__isNat (x11:Z) := 0. Definition P_id_U43 (x11:Z) := 2* x11. Definition P_id_a__isNatIListKind (x11:Z) := 0. Definition P_id_a__U12 (x11:Z) := 1* x11. Definition P_id_U31 (x11:Z) (x12:Z) := 2* x11. Definition P_id_a__length (x11:Z) := 1 + 1* x11. Definition P_id_a__U42 (x11:Z) (x12:Z) := 1* x11. Definition P_id_U53 (x11:Z) := 1* x11. Definition P_id_and (x11:Z) (x12:Z) := 2* x11 + 1* x12. Definition P_id_cons (x11:Z) (x12:Z) := 2* x11 + 2* x12. Definition P_id_U11 (x11:Z) (x12:Z) := 1* x11. Definition P_id_a__U51 (x11:Z) (x12:Z) (x13:Z) := 1* x11. Definition P_id_a__U22 (x11:Z) := 2* x11. Definition P_id_U42 (x11:Z) (x12:Z) := 1* x11. Definition P_id_length (x11:Z) := 1 + 1* x11. Definition P_id_tt := 0. Definition P_id_U22 (x11:Z) := 2* x11. Definition P_id_s (x11:Z) := 1* x11. Definition P_id_a__U41 (x11:Z) (x12:Z) (x13:Z) := 2* x11. Definition P_id_U52 (x11:Z) (x12:Z) := 2* x11. Definition P_id_nil := 3. Definition P_id_zeros := 0. Definition P_id_isNatList (x11:Z) := 0. Definition P_id_a__U53 (x11:Z) := 1* x11. Definition P_id_a__U31 (x11:Z) (x12:Z) := 2* x11. Definition P_id_isNatIList (x11:Z) := 0. Definition P_id_a__isNatKind (x11:Z) := 0. Definition P_id_a__isNatList (x11:Z) := 0. Definition P_id_U32 (x11:Z) := 2* x11. Definition P_id_mark (x11:Z) := 1* x11. Definition P_id_a__U43 (x11:Z) := 2* x11. Definition P_id_U61 (x11:Z) (x12:Z) := 1 + 2* x11 + 1* x12. Definition P_id_isNat (x11:Z) := 0. Lemma P_id_isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNatKind x12 <= P_id_isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIList x12 <= P_id_a__isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_a__U21 x12 x14 <= P_id_a__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U41 x12 x14 x16 <= P_id_U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_a__and x12 x14 <= P_id_a__and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_a__U11 x12 x14 <= P_id_a__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U21 x12 x14 <= P_id_U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_a__U61 x12 x14 <= P_id_a__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U32 x12 <= P_id_a__U32 x11. Proof. intros x12 x11. intros [H_1 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 x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_U51 x12 x14 x16 <= P_id_U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_isNatIListKind x12 <= P_id_isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U12 x12 <= P_id_U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_a__U52 x12 x14 <= P_id_a__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__isNat x12 <= P_id_a__isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U43 x12 <= P_id_U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatIListKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatIListKind x12 <= P_id_a__isNatIListKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U12 x12 <= P_id_a__U12 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U31 x12 x14 <= P_id_U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__length x12 <= P_id_a__length x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_a__U42 x12 x14 <= P_id_a__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U53 x12 <= P_id_U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_and x12 x14 <= P_id_and x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_cons x12 x14 <= P_id_cons x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 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_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U11 x12 x14 <= P_id_U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U51 x12 x14 x16 <= P_id_a__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U22 x12 <= P_id_a__U22 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U42 x12 x14 <= P_id_U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_length x12 <= P_id_length x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U22 x12 <= P_id_U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_s x12 <= P_id_s x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_a__U41 x12 x14 x16 <= P_id_a__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U52 x12 x14 <= P_id_U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNatList x12 <= P_id_isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U53 x12 <= P_id_a__U53 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_a__U31 x12 x14 <= P_id_a__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNatIList x12 <= P_id_isNatIList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatKind_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatKind x12 <= P_id_a__isNatKind x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNatList_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_a__isNatList x12 <= P_id_a__isNatList x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_U32 x12 <= P_id_U32 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_mark x12 <= P_id_mark x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_a__U43 x12 <= P_id_a__U43 x11. Proof. intros x12 x11. intros [H_1 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 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_U61 x12 x14 <= P_id_U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_isNat x12 <= P_id_isNat x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__zeros_bounded : 0 <= P_id_a__zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatKind x11. 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__isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIList x11. 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__U21_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U21 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_U41 x13 x12 x11. 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__and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__and x12 x11. 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__U11_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U11 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U21 x12 x11. 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__U61_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U61 x12 x11. 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__U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U32 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_U51 x13 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIListKind x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_U12 x11. 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__U52_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U52 x12 x11. 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__isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNat x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_U43 x11. 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__isNatIListKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatIListKind x11. 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__U12_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U31 x12 x11. 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__length_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__length x11. 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__U42_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_U53 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_and x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_cons x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U11 x12 x11. 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__U51_bounded : forall x12 x13 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U51 x13 x12 x11. 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__U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U22 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U42 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_length_bounded : forall x11, (0 <= x11) ->0 <= P_id_length x11. 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_U22_bounded : forall x11, (0 <= x11) ->0 <= P_id_U22 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x11, (0 <= x11) ->0 <= P_id_s x11. 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__U41_bounded : forall x12 x13 x11, (0 <= x11) ->(0 <= x12) ->(0 <= x13) ->0 <= P_id_a__U41 x13 x12 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U52 x12 x11. 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_zeros_bounded : 0 <= P_id_zeros . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatList x11. 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__U53_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U53 x11. 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__U31_bounded : forall x12 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_a__U31 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNatIList_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNatIList x11. 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__isNatKind_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatKind x11. 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__isNatList_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__isNatList x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U32_bounded : forall x11, (0 <= x11) ->0 <= P_id_U32 x11. 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 x11, (0 <= x11) ->0 <= P_id_mark x11. 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__U43_bounded : forall x11, (0 <= x11) ->0 <= P_id_a__U43 x11. 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 x11, (0 <= x11) ->(0 <= x12) ->0 <= P_id_U61 x12 x11. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNat_bounded : forall x11, (0 <= x11) ->0 <= P_id_isNat x11. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_a__zeros | (algebra.Alg.Term algebra.F.id_isNatKind (x11::nil)) => P_id_isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_a__isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12::x11::nil)) => P_id_a__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U41 (x13::x12::x11::nil)) => P_id_U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12::x11::nil)) => P_id_a__and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12::x11::nil)) => P_id_a__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U21 (x12::x11::nil)) => P_id_U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12::x11::nil)) => P_id_a__U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_a__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_U51 (x13::x12::x11::nil)) => P_id_U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIListKind (x11::nil)) => P_id_isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_U12 (x11::nil)) => P_id_U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12::x11::nil)) => P_id_a__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_a__isNat (measure x11) | (algebra.Alg.Term algebra.F.id_U43 (x11::nil)) => P_id_U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_a__isNatIListKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_a__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_U31 (x12::x11::nil)) => P_id_U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_a__length (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12::x11::nil)) => P_id_a__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U53 (x11::nil)) => P_id_U53 (measure x11) | (algebra.Alg.Term algebra.F.id_and (x12::x11::nil)) => P_id_and (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_cons (x12::x11::nil)) => P_id_cons (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U11 (x12::x11::nil)) => P_id_U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_a__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_a__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_U42 (x12::x11::nil)) => P_id_U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_length (x11::nil)) => P_id_length (measure x11) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U22 (x11::nil)) => P_id_U22 (measure x11) | (algebra.Alg.Term algebra.F.id_s (x11::nil)) => P_id_s (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_a__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_U52 (x12::x11::nil)) => P_id_U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_zeros nil) => P_id_zeros | (algebra.Alg.Term algebra.F.id_isNatList (x11::nil)) => P_id_isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_a__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12::x11::nil)) => P_id_a__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNatIList (x11::nil)) => P_id_isNatIList (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_a__isNatKind (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_a__isNatList (measure x11) | (algebra.Alg.Term algebra.F.id_U32 (x11::nil)) => P_id_U32 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_mark (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_a__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_U61 (x12::x11::nil)) => P_id_U61 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_isNat (x11::nil)) => P_id_isNat (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__LENGTH (x11:Z) := 1 + 1* x11. Definition P_id_A__U11 (x11:Z) (x12:Z) := 0. Definition P_id_A__U41 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__ISNATKIND (x11:Z) := 0. Definition P_id_A__U22 (x11:Z) := 0. Definition P_id_A__U51 (x11:Z) (x12:Z) (x13:Z) := 0. Definition P_id_A__AND (x11:Z) (x12:Z) := 2* x12. Definition P_id_A__ISNATLIST (x11:Z) := 0. Definition P_id_A__U43 (x11:Z) := 0. Definition P_id_A__U31 (x11:Z) (x12:Z) := 0. Definition P_id_A__U53 (x11:Z) := 0. Definition P_id_MARK (x11:Z) := 2* x11. Definition P_id_A__U12 (x11:Z) := 0. Definition P_id_A__U42 (x11:Z) (x12:Z) := 0. Definition P_id_A__ZEROS := 0. Definition P_id_A__ISNAT (x11:Z) := 0. Definition P_id_A__U52 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILISTKIND (x11:Z) := 0. Definition P_id_A__U21 (x11:Z) (x12:Z) := 0. Definition P_id_A__ISNATILIST (x11:Z) := 0. Definition P_id_A__U32 (x11:Z) := 0. Definition P_id_A__U61 (x11:Z) (x12:Z) := 1 + 2* x12. Lemma P_id_A__LENGTH_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__LENGTH x12 <= P_id_A__LENGTH x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U11_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_A__U11 x12 x14 <= P_id_A__U11 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U41_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U41 x12 x14 x16 <= P_id_A__U41 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATKIND x12 <= P_id_A__ISNATKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U22_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U22 x12 <= P_id_A__U22 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U51_monotonic : forall x16 x12 x14 x13 x11 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) -> P_id_A__U51 x12 x14 x16 <= P_id_A__U51 x11 x13 x15. Proof. intros x16 x15 x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__AND_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_A__AND x12 x14 <= P_id_A__AND x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATLIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATLIST x12 <= P_id_A__ISNATLIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U43_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U43 x12 <= P_id_A__U43 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U31_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_A__U31 x12 x14 <= P_id_A__U31 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U53_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U53 x12 <= P_id_A__U53 x11. Proof. intros x12 x11. intros [H_1 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 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_MARK x12 <= P_id_MARK x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U12_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U12 x12 <= P_id_A__U12 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U42_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_A__U42 x12 x14 <= P_id_A__U42 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNAT_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__ISNAT x12 <= P_id_A__ISNAT x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U52_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_A__U52 x12 x14 <= P_id_A__U52 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILISTKIND_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILISTKIND x12 <= P_id_A__ISNATILISTKIND x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U21_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_A__U21 x12 x14 <= P_id_A__U21 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISNATILIST_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) -> P_id_A__ISNATILIST x12 <= P_id_A__ISNATILIST x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U32_monotonic : forall x12 x11, (0 <= x12)/\ (x12 <= x11) ->P_id_A__U32 x12 <= P_id_A__U32 x11. Proof. intros x12 x11. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U61_monotonic : forall x12 x14 x13 x11, (0 <= x14)/\ (x14 <= x13) -> (0 <= x12)/\ (x12 <= x11) ->P_id_A__U61 x12 x14 <= P_id_A__U61 x11 x13. Proof. intros x14 x13 x12 x11. intros [H_1 H_0]. 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_a__zeros P_id_isNatKind P_id_a__isNatIList P_id_a__U21 P_id_U41 P_id_a__and P_id_a__U11 P_id_U21 P_id_a__U61 P_id_a__U32 P_id_U51 P_id_isNatIListKind P_id_0 P_id_U12 P_id_a__U52 P_id_a__isNat P_id_U43 P_id_a__isNatIListKind P_id_a__U12 P_id_U31 P_id_a__length P_id_a__U42 P_id_U53 P_id_and P_id_cons P_id_U11 P_id_a__U51 P_id_a__U22 P_id_U42 P_id_length P_id_tt P_id_U22 P_id_s P_id_a__U41 P_id_U52 P_id_nil P_id_zeros P_id_isNatList P_id_a__U53 P_id_a__U31 P_id_isNatIList P_id_a__isNatKind P_id_a__isNatList P_id_U32 P_id_mark P_id_a__U43 P_id_U61 P_id_isNat P_id_A__LENGTH P_id_A__U11 P_id_A__U41 P_id_A__ISNATKIND P_id_A__U22 P_id_A__U51 P_id_A__AND P_id_A__ISNATLIST P_id_A__U43 P_id_A__U31 P_id_A__U53 P_id_MARK P_id_A__U12 P_id_A__U42 P_id_A__ZEROS P_id_A__ISNAT P_id_A__U52 P_id_A__ISNATILISTKIND P_id_A__U21 P_id_A__ISNATILIST P_id_A__U32 P_id_A__U61. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__length (x11::nil)) => P_id_A__LENGTH (measure x11) | (algebra.Alg.Term algebra.F.id_a__U11 (x12:: x11::nil)) => P_id_A__U11 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U41 (x13::x12:: x11::nil)) => P_id_A__U41 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatKind (x11::nil)) => P_id_A__ISNATKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U22 (x11::nil)) => P_id_A__U22 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U51 (x13::x12:: x11::nil)) => P_id_A__U51 (measure x13) (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__and (x12:: x11::nil)) => P_id_A__AND (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatList (x11::nil)) => P_id_A__ISNATLIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U43 (x11::nil)) => P_id_A__U43 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U31 (x12:: x11::nil)) => P_id_A__U31 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__U53 (x11::nil)) => P_id_A__U53 (measure x11) | (algebra.Alg.Term algebra.F.id_mark (x11::nil)) => P_id_MARK (measure x11) | (algebra.Alg.Term algebra.F.id_a__U12 (x11::nil)) => P_id_A__U12 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U42 (x12:: x11::nil)) => P_id_A__U42 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__zeros nil) => P_id_A__ZEROS | (algebra.Alg.Term algebra.F.id_a__isNat (x11::nil)) => P_id_A__ISNAT (measure x11) | (algebra.Alg.Term algebra.F.id_a__U52 (x12:: x11::nil)) => P_id_A__U52 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIListKind (x11::nil)) => P_id_A__ISNATILISTKIND (measure x11) | (algebra.Alg.Term algebra.F.id_a__U21 (x12:: x11::nil)) => P_id_A__U21 (measure x12) (measure x11) | (algebra.Alg.Term algebra.F.id_a__isNatIList (x11::nil)) => P_id_A__ISNATILIST (measure x11) | (algebra.Alg.Term algebra.F.id_a__U32 (x11::nil)) => P_id_A__U32 (measure x11) | (algebra.Alg.Term algebra.F.id_a__U61 (x12:: x11::nil)) => P_id_A__U61 (measure x12) (measure x11) | _ => 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_isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatIList_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_isNatIListKind_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_a__isNat_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__isNatIListKind_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_a__length_monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_length_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_isNatList_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_isNatIList_monotonic;assumption. intros ;apply P_id_a__isNatKind_monotonic;assumption. intros ;apply P_id_a__isNatList_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_isNat_monotonic;assumption. intros ;apply P_id_a__zeros_bounded;assumption. intros ;apply P_id_isNatKind_bounded;assumption. intros ;apply P_id_a__isNatIList_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_isNatIListKind_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_a__isNat_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__isNatIListKind_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_a__length_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_length_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_zeros_bounded;assumption. intros ;apply P_id_isNatList_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_isNatIList_bounded;assumption. intros ;apply P_id_a__isNatKind_bounded;assumption. intros ;apply P_id_a__isNatList_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_isNat_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__LENGTH_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__ISNATKIND_monotonic;assumption. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__ISNATLIST_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__ISNAT_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A__ISNATILISTKIND_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. intros ;apply P_id_A__ISNATILIST_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_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_12_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_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_12_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_12_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_12_large := WF_DP_R_xml_0_scc_12_large.wf. Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_12. Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx';apply DP_R_xml_0_scc_12_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_12_large_in_le;econstructor eassumption])). apply wf_DP_R_xml_0_scc_12_large. Qed. End WF_DP_R_xml_0_scc_12. Definition wf_DP_R_xml_0_scc_12 := WF_DP_R_xml_0_scc_12.wf. Lemma acc_DP_R_xml_0_scc_12 : forall x y, (DP_R_xml_0_scc_12 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_12). 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_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' ))|| ((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_12. 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_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_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___LengthOfFiniteLists_complete_GM.trs/a3pat.v" *** *** End: *** *)