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____ *) | id_a____ : symb (* id_a *) | id_a : symb (* id_a__U42 *) | id_a__U42 : symb (* id_U42 *) | id_U42 : symb (* id_a__U21 *) | id_a__U21 : symb (* id_U21 *) | id_U21 : symb (* id_a__U72 *) | id_a__U72 : symb (* id_U72 *) | id_U72 : symb (* id_a__U11 *) | id_a__U11 : symb (* id_u *) | id_u : symb (* id_a__U53 *) | id_a__U53 : symb (* id_U53 *) | id_U53 : symb (* id_a__U31 *) | id_a__U31 : symb (* id_U31 *) | id_U31 : symb (* id_isPalListKind *) | id_isPalListKind : symb (* id_mark *) | id_mark : symb (* id_i *) | id_i : symb (* id_a__U51 *) | id_a__U51 : symb (* id_U51 *) | id_U51 : symb (* id_a__isList *) | id_a__isList : symb (* id_isList *) | id_isList : symb (* id_a__and *) | id_a__and : symb (* id_a__U12 *) | id_a__U12 : symb (* id_U12 *) | id_U12 : symb (* id_a__U62 *) | id_a__U62 : symb (* id_U62 *) | id_U62 : symb (* id_a__isQid *) | id_a__isQid : symb (* id_isQid *) | id_isQid : symb (* id_isPal *) | id_isPal : symb (* id___ *) | id___ : symb (* id_e *) | id_e : symb (* id_a__U43 *) | id_a__U43 : symb (* id_U43 *) | id_U43 : symb (* id_a__U22 *) | id_a__U22 : symb (* id_U22 *) | id_U22 : symb (* id_a__isNePal *) | id_a__isNePal : symb (* id_isNePal *) | id_isNePal : symb (* id_tt *) | id_tt : symb (* id_U11 *) | id_U11 : symb (* id_a__U61 *) | id_a__U61 : symb (* id_U61 *) | id_U61 : symb (* id_a__U32 *) | id_a__U32 : symb (* id_U32 *) | id_U32 : symb (* id_and *) | id_and : symb (* id_nil *) | id_nil : symb (* id_o *) | id_o : symb (* id_a__U52 *) | id_a__U52 : symb (* id_U52 *) | id_U52 : symb (* id_a__U23 *) | id_a__U23 : symb (* id_U23 *) | id_U23 : symb (* id_a__isPalListKind *) | id_a__isPalListKind : symb (* id_a__isNeList *) | id_a__isNeList : symb (* id_isNeList *) | id_isNeList : symb (* id_a__U71 *) | id_a__U71 : symb (* id_U71 *) | id_U71 : symb (* id_a__U41 *) | id_a__U41 : symb (* id_U41 *) | id_U41 : symb (* id_a__isPal *) | id_a__isPal : symb . Definition symb_eq_bool (f1 f2:symb) : bool := match f1,f2 with | id_a____,id_a____ => true | id_a,id_a => true | id_a__U42,id_a__U42 => true | id_U42,id_U42 => true | id_a__U21,id_a__U21 => true | id_U21,id_U21 => true | id_a__U72,id_a__U72 => true | id_U72,id_U72 => true | id_a__U11,id_a__U11 => true | id_u,id_u => true | id_a__U53,id_a__U53 => true | id_U53,id_U53 => true | id_a__U31,id_a__U31 => true | id_U31,id_U31 => true | id_isPalListKind,id_isPalListKind => true | id_mark,id_mark => true | id_i,id_i => true | id_a__U51,id_a__U51 => true | id_U51,id_U51 => true | id_a__isList,id_a__isList => true | id_isList,id_isList => true | id_a__and,id_a__and => true | id_a__U12,id_a__U12 => true | id_U12,id_U12 => true | id_a__U62,id_a__U62 => true | id_U62,id_U62 => true | id_a__isQid,id_a__isQid => true | id_isQid,id_isQid => true | id_isPal,id_isPal => true | id___,id___ => true | id_e,id_e => true | id_a__U43,id_a__U43 => true | id_U43,id_U43 => true | id_a__U22,id_a__U22 => true | id_U22,id_U22 => true | id_a__isNePal,id_a__isNePal => true | id_isNePal,id_isNePal => true | id_tt,id_tt => true | id_U11,id_U11 => true | id_a__U61,id_a__U61 => true | id_U61,id_U61 => true | id_a__U32,id_a__U32 => true | id_U32,id_U32 => true | id_and,id_and => true | id_nil,id_nil => true | id_o,id_o => true | id_a__U52,id_a__U52 => true | id_U52,id_U52 => true | id_a__U23,id_a__U23 => true | id_U23,id_U23 => true | id_a__isPalListKind,id_a__isPalListKind => true | id_a__isNeList,id_a__isNeList => true | id_isNeList,id_isNeList => true | id_a__U71,id_a__U71 => true | id_U71,id_U71 => true | id_a__U41,id_a__U41 => true | id_U41,id_U41 => true | id_a__isPal,id_a__isPal => 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____,id_a____ => refl_equal _ | id_a,id_a => refl_equal _ | id_a__U42,id_a__U42 => refl_equal _ | id_U42,id_U42 => refl_equal _ | id_a__U21,id_a__U21 => refl_equal _ | id_U21,id_U21 => refl_equal _ | id_a__U72,id_a__U72 => refl_equal _ | id_U72,id_U72 => refl_equal _ | id_a__U11,id_a__U11 => refl_equal _ | id_u,id_u => refl_equal _ | id_a__U53,id_a__U53 => refl_equal _ | id_U53,id_U53 => refl_equal _ | id_a__U31,id_a__U31 => refl_equal _ | id_U31,id_U31 => refl_equal _ | id_isPalListKind,id_isPalListKind => refl_equal _ | id_mark,id_mark => refl_equal _ | id_i,id_i => refl_equal _ | id_a__U51,id_a__U51 => refl_equal _ | id_U51,id_U51 => refl_equal _ | id_a__isList,id_a__isList => refl_equal _ | id_isList,id_isList => refl_equal _ | id_a__and,id_a__and => refl_equal _ | id_a__U12,id_a__U12 => refl_equal _ | id_U12,id_U12 => refl_equal _ | id_a__U62,id_a__U62 => refl_equal _ | id_U62,id_U62 => refl_equal _ | id_a__isQid,id_a__isQid => refl_equal _ | id_isQid,id_isQid => refl_equal _ | id_isPal,id_isPal => refl_equal _ | id___,id___ => refl_equal _ | id_e,id_e => refl_equal _ | id_a__U43,id_a__U43 => refl_equal _ | id_U43,id_U43 => refl_equal _ | id_a__U22,id_a__U22 => refl_equal _ | id_U22,id_U22 => refl_equal _ | id_a__isNePal,id_a__isNePal => refl_equal _ | id_isNePal,id_isNePal => refl_equal _ | id_tt,id_tt => refl_equal _ | id_U11,id_U11 => refl_equal _ | id_a__U61,id_a__U61 => refl_equal _ | id_U61,id_U61 => refl_equal _ | id_a__U32,id_a__U32 => refl_equal _ | id_U32,id_U32 => refl_equal _ | id_and,id_and => refl_equal _ | id_nil,id_nil => refl_equal _ | id_o,id_o => refl_equal _ | id_a__U52,id_a__U52 => refl_equal _ | id_U52,id_U52 => refl_equal _ | id_a__U23,id_a__U23 => refl_equal _ | id_U23,id_U23 => refl_equal _ | id_a__isPalListKind,id_a__isPalListKind => refl_equal _ | id_a__isNeList,id_a__isNeList => refl_equal _ | id_isNeList,id_isNeList => refl_equal _ | id_a__U71,id_a__U71 => refl_equal _ | id_U71,id_U71 => refl_equal _ | id_a__U41,id_a__U41 => refl_equal _ | id_U41,id_U41 => refl_equal _ | id_a__isPal,id_a__isPal => refl_equal _ | _,_ => _ end;intros abs;discriminate. Defined. Definition arity (f:symb) := match f with | id_a____ => term.Free 2 | id_a => term.Free 0 | id_a__U42 => term.Free 2 | id_U42 => term.Free 2 | id_a__U21 => term.Free 3 | id_U21 => term.Free 3 | id_a__U72 => term.Free 1 | id_U72 => term.Free 1 | id_a__U11 => term.Free 2 | id_u => term.Free 0 | id_a__U53 => term.Free 1 | id_U53 => term.Free 1 | id_a__U31 => term.Free 2 | id_U31 => term.Free 2 | id_isPalListKind => term.Free 1 | id_mark => term.Free 1 | id_i => term.Free 0 | id_a__U51 => term.Free 3 | id_U51 => term.Free 3 | id_a__isList => term.Free 1 | id_isList => term.Free 1 | id_a__and => term.Free 2 | id_a__U12 => term.Free 1 | id_U12 => term.Free 1 | id_a__U62 => term.Free 1 | id_U62 => term.Free 1 | id_a__isQid => term.Free 1 | id_isQid => term.Free 1 | id_isPal => term.Free 1 | id___ => term.Free 2 | id_e => term.Free 0 | id_a__U43 => term.Free 1 | id_U43 => term.Free 1 | id_a__U22 => term.Free 2 | id_U22 => term.Free 2 | id_a__isNePal => term.Free 1 | id_isNePal => term.Free 1 | id_tt => term.Free 0 | id_U11 => term.Free 2 | id_a__U61 => term.Free 2 | id_U61 => term.Free 2 | id_a__U32 => term.Free 1 | id_U32 => term.Free 1 | id_and => term.Free 2 | id_nil => term.Free 0 | id_o => term.Free 0 | id_a__U52 => term.Free 2 | id_U52 => term.Free 2 | id_a__U23 => term.Free 1 | id_U23 => term.Free 1 | id_a__isPalListKind => term.Free 1 | id_a__isNeList => term.Free 1 | id_isNeList => term.Free 1 | id_a__U71 => term.Free 2 | id_U71 => term.Free 2 | id_a__U41 => term.Free 3 | id_U41 => term.Free 3 | id_a__isPal => term.Free 1 end. Definition symb_order (f1 f2:symb) : bool := match f1,f2 with | id_a____,id_a____ => true | id_a____,id_a => false | id_a____,id_a__U42 => false | id_a____,id_U42 => false | id_a____,id_a__U21 => false | id_a____,id_U21 => false | id_a____,id_a__U72 => false | id_a____,id_U72 => false | id_a____,id_a__U11 => false | id_a____,id_u => false | id_a____,id_a__U53 => false | id_a____,id_U53 => false | id_a____,id_a__U31 => false | id_a____,id_U31 => false | id_a____,id_isPalListKind => false | id_a____,id_mark => false | id_a____,id_i => false | id_a____,id_a__U51 => false | id_a____,id_U51 => false | id_a____,id_a__isList => false | id_a____,id_isList => false | id_a____,id_a__and => false | id_a____,id_a__U12 => false | id_a____,id_U12 => false | id_a____,id_a__U62 => false | id_a____,id_U62 => false | id_a____,id_a__isQid => false | id_a____,id_isQid => false | id_a____,id_isPal => false | id_a____,id___ => false | id_a____,id_e => false | id_a____,id_a__U43 => false | id_a____,id_U43 => false | id_a____,id_a__U22 => false | id_a____,id_U22 => false | id_a____,id_a__isNePal => false | id_a____,id_isNePal => false | id_a____,id_tt => false | id_a____,id_U11 => false | id_a____,id_a__U61 => false | id_a____,id_U61 => false | id_a____,id_a__U32 => false | id_a____,id_U32 => false | id_a____,id_and => false | id_a____,id_nil => false | id_a____,id_o => false | id_a____,id_a__U52 => false | id_a____,id_U52 => false | id_a____,id_a__U23 => false | id_a____,id_U23 => false | id_a____,id_a__isPalListKind => false | id_a____,id_a__isNeList => false | id_a____,id_isNeList => false | id_a____,id_a__U71 => false | id_a____,id_U71 => false | id_a____,id_a__U41 => false | id_a____,id_U41 => false | id_a____,id_a__isPal => false | id_a,id_a____ => true | id_a,id_a => true | id_a,id_a__U42 => false | id_a,id_U42 => false | id_a,id_a__U21 => false | id_a,id_U21 => false | id_a,id_a__U72 => false | id_a,id_U72 => false | id_a,id_a__U11 => false | id_a,id_u => false | id_a,id_a__U53 => false | id_a,id_U53 => false | id_a,id_a__U31 => false | id_a,id_U31 => false | id_a,id_isPalListKind => false | id_a,id_mark => false | id_a,id_i => false | id_a,id_a__U51 => false | id_a,id_U51 => false | id_a,id_a__isList => false | id_a,id_isList => false | id_a,id_a__and => false | id_a,id_a__U12 => false | id_a,id_U12 => false | id_a,id_a__U62 => false | id_a,id_U62 => false | id_a,id_a__isQid => false | id_a,id_isQid => false | id_a,id_isPal => false | id_a,id___ => false | id_a,id_e => false | id_a,id_a__U43 => false | id_a,id_U43 => false | id_a,id_a__U22 => false | id_a,id_U22 => false | id_a,id_a__isNePal => false | id_a,id_isNePal => false | id_a,id_tt => false | id_a,id_U11 => false | id_a,id_a__U61 => false | id_a,id_U61 => false | id_a,id_a__U32 => false | id_a,id_U32 => false | id_a,id_and => false | id_a,id_nil => false | id_a,id_o => false | id_a,id_a__U52 => false | id_a,id_U52 => false | id_a,id_a__U23 => false | id_a,id_U23 => false | id_a,id_a__isPalListKind => false | id_a,id_a__isNeList => false | id_a,id_isNeList => false | id_a,id_a__U71 => false | id_a,id_U71 => false | id_a,id_a__U41 => false | id_a,id_U41 => false | id_a,id_a__isPal => false | id_a__U42,id_a____ => true | id_a__U42,id_a => true | id_a__U42,id_a__U42 => true | id_a__U42,id_U42 => false | id_a__U42,id_a__U21 => false | id_a__U42,id_U21 => false | id_a__U42,id_a__U72 => false | id_a__U42,id_U72 => false | id_a__U42,id_a__U11 => false | id_a__U42,id_u => false | id_a__U42,id_a__U53 => false | id_a__U42,id_U53 => false | id_a__U42,id_a__U31 => false | id_a__U42,id_U31 => false | id_a__U42,id_isPalListKind => false | id_a__U42,id_mark => false | id_a__U42,id_i => false | id_a__U42,id_a__U51 => false | id_a__U42,id_U51 => false | id_a__U42,id_a__isList => false | id_a__U42,id_isList => false | id_a__U42,id_a__and => false | id_a__U42,id_a__U12 => false | id_a__U42,id_U12 => false | id_a__U42,id_a__U62 => false | id_a__U42,id_U62 => false | id_a__U42,id_a__isQid => false | id_a__U42,id_isQid => false | id_a__U42,id_isPal => false | id_a__U42,id___ => false | id_a__U42,id_e => false | id_a__U42,id_a__U43 => false | id_a__U42,id_U43 => false | id_a__U42,id_a__U22 => false | id_a__U42,id_U22 => false | id_a__U42,id_a__isNePal => false | id_a__U42,id_isNePal => false | id_a__U42,id_tt => false | id_a__U42,id_U11 => false | id_a__U42,id_a__U61 => false | id_a__U42,id_U61 => false | id_a__U42,id_a__U32 => false | id_a__U42,id_U32 => false | id_a__U42,id_and => false | id_a__U42,id_nil => false | id_a__U42,id_o => false | id_a__U42,id_a__U52 => false | id_a__U42,id_U52 => false | id_a__U42,id_a__U23 => false | id_a__U42,id_U23 => false | id_a__U42,id_a__isPalListKind => false | id_a__U42,id_a__isNeList => false | id_a__U42,id_isNeList => false | id_a__U42,id_a__U71 => false | id_a__U42,id_U71 => false | id_a__U42,id_a__U41 => false | id_a__U42,id_U41 => false | id_a__U42,id_a__isPal => false | id_U42,id_a____ => true | id_U42,id_a => true | id_U42,id_a__U42 => true | id_U42,id_U42 => true | id_U42,id_a__U21 => false | id_U42,id_U21 => false | id_U42,id_a__U72 => false | id_U42,id_U72 => false | id_U42,id_a__U11 => false | id_U42,id_u => false | id_U42,id_a__U53 => false | id_U42,id_U53 => false | id_U42,id_a__U31 => false | id_U42,id_U31 => false | id_U42,id_isPalListKind => false | id_U42,id_mark => false | id_U42,id_i => false | id_U42,id_a__U51 => false | id_U42,id_U51 => false | id_U42,id_a__isList => false | id_U42,id_isList => false | id_U42,id_a__and => false | id_U42,id_a__U12 => false | id_U42,id_U12 => false | id_U42,id_a__U62 => false | id_U42,id_U62 => false | id_U42,id_a__isQid => false | id_U42,id_isQid => false | id_U42,id_isPal => false | id_U42,id___ => false | id_U42,id_e => false | id_U42,id_a__U43 => false | id_U42,id_U43 => false | id_U42,id_a__U22 => false | id_U42,id_U22 => false | id_U42,id_a__isNePal => false | id_U42,id_isNePal => false | id_U42,id_tt => false | id_U42,id_U11 => false | id_U42,id_a__U61 => false | id_U42,id_U61 => false | id_U42,id_a__U32 => false | id_U42,id_U32 => false | id_U42,id_and => false | id_U42,id_nil => false | id_U42,id_o => false | id_U42,id_a__U52 => false | id_U42,id_U52 => false | id_U42,id_a__U23 => false | id_U42,id_U23 => false | id_U42,id_a__isPalListKind => false | id_U42,id_a__isNeList => false | id_U42,id_isNeList => false | id_U42,id_a__U71 => false | id_U42,id_U71 => false | id_U42,id_a__U41 => false | id_U42,id_U41 => false | id_U42,id_a__isPal => false | id_a__U21,id_a____ => true | id_a__U21,id_a => true | id_a__U21,id_a__U42 => true | id_a__U21,id_U42 => true | id_a__U21,id_a__U21 => true | id_a__U21,id_U21 => false | id_a__U21,id_a__U72 => false | id_a__U21,id_U72 => false | id_a__U21,id_a__U11 => false | id_a__U21,id_u => false | id_a__U21,id_a__U53 => false | id_a__U21,id_U53 => false | id_a__U21,id_a__U31 => false | id_a__U21,id_U31 => false | id_a__U21,id_isPalListKind => false | id_a__U21,id_mark => false | id_a__U21,id_i => false | id_a__U21,id_a__U51 => false | id_a__U21,id_U51 => false | id_a__U21,id_a__isList => false | id_a__U21,id_isList => false | id_a__U21,id_a__and => false | id_a__U21,id_a__U12 => false | id_a__U21,id_U12 => false | id_a__U21,id_a__U62 => false | id_a__U21,id_U62 => false | id_a__U21,id_a__isQid => false | id_a__U21,id_isQid => false | id_a__U21,id_isPal => false | id_a__U21,id___ => false | id_a__U21,id_e => false | id_a__U21,id_a__U43 => false | id_a__U21,id_U43 => false | id_a__U21,id_a__U22 => false | id_a__U21,id_U22 => false | id_a__U21,id_a__isNePal => false | id_a__U21,id_isNePal => false | id_a__U21,id_tt => false | id_a__U21,id_U11 => false | id_a__U21,id_a__U61 => false | id_a__U21,id_U61 => false | id_a__U21,id_a__U32 => false | id_a__U21,id_U32 => false | id_a__U21,id_and => false | id_a__U21,id_nil => false | id_a__U21,id_o => false | id_a__U21,id_a__U52 => false | id_a__U21,id_U52 => false | id_a__U21,id_a__U23 => false | id_a__U21,id_U23 => false | id_a__U21,id_a__isPalListKind => false | id_a__U21,id_a__isNeList => false | id_a__U21,id_isNeList => false | id_a__U21,id_a__U71 => false | id_a__U21,id_U71 => false | id_a__U21,id_a__U41 => false | id_a__U21,id_U41 => false | id_a__U21,id_a__isPal => false | id_U21,id_a____ => true | id_U21,id_a => true | id_U21,id_a__U42 => true | id_U21,id_U42 => true | id_U21,id_a__U21 => true | id_U21,id_U21 => true | id_U21,id_a__U72 => false | id_U21,id_U72 => false | id_U21,id_a__U11 => false | id_U21,id_u => false | id_U21,id_a__U53 => false | id_U21,id_U53 => false | id_U21,id_a__U31 => false | id_U21,id_U31 => false | id_U21,id_isPalListKind => false | id_U21,id_mark => false | id_U21,id_i => false | id_U21,id_a__U51 => false | id_U21,id_U51 => false | id_U21,id_a__isList => false | id_U21,id_isList => false | id_U21,id_a__and => false | id_U21,id_a__U12 => false | id_U21,id_U12 => false | id_U21,id_a__U62 => false | id_U21,id_U62 => false | id_U21,id_a__isQid => false | id_U21,id_isQid => false | id_U21,id_isPal => false | id_U21,id___ => false | id_U21,id_e => false | id_U21,id_a__U43 => false | id_U21,id_U43 => false | id_U21,id_a__U22 => false | id_U21,id_U22 => false | id_U21,id_a__isNePal => false | id_U21,id_isNePal => false | id_U21,id_tt => false | id_U21,id_U11 => false | id_U21,id_a__U61 => false | id_U21,id_U61 => false | id_U21,id_a__U32 => false | id_U21,id_U32 => false | id_U21,id_and => false | id_U21,id_nil => false | id_U21,id_o => false | id_U21,id_a__U52 => false | id_U21,id_U52 => false | id_U21,id_a__U23 => false | id_U21,id_U23 => false | id_U21,id_a__isPalListKind => false | id_U21,id_a__isNeList => false | id_U21,id_isNeList => false | id_U21,id_a__U71 => false | id_U21,id_U71 => false | id_U21,id_a__U41 => false | id_U21,id_U41 => false | id_U21,id_a__isPal => false | id_a__U72,id_a____ => true | id_a__U72,id_a => true | id_a__U72,id_a__U42 => true | id_a__U72,id_U42 => true | id_a__U72,id_a__U21 => true | id_a__U72,id_U21 => true | id_a__U72,id_a__U72 => true | id_a__U72,id_U72 => false | id_a__U72,id_a__U11 => false | id_a__U72,id_u => false | id_a__U72,id_a__U53 => false | id_a__U72,id_U53 => false | id_a__U72,id_a__U31 => false | id_a__U72,id_U31 => false | id_a__U72,id_isPalListKind => false | id_a__U72,id_mark => false | id_a__U72,id_i => false | id_a__U72,id_a__U51 => false | id_a__U72,id_U51 => false | id_a__U72,id_a__isList => false | id_a__U72,id_isList => false | id_a__U72,id_a__and => false | id_a__U72,id_a__U12 => false | id_a__U72,id_U12 => false | id_a__U72,id_a__U62 => false | id_a__U72,id_U62 => false | id_a__U72,id_a__isQid => false | id_a__U72,id_isQid => false | id_a__U72,id_isPal => false | id_a__U72,id___ => false | id_a__U72,id_e => false | id_a__U72,id_a__U43 => false | id_a__U72,id_U43 => false | id_a__U72,id_a__U22 => false | id_a__U72,id_U22 => false | id_a__U72,id_a__isNePal => false | id_a__U72,id_isNePal => false | id_a__U72,id_tt => false | id_a__U72,id_U11 => false | id_a__U72,id_a__U61 => false | id_a__U72,id_U61 => false | id_a__U72,id_a__U32 => false | id_a__U72,id_U32 => false | id_a__U72,id_and => false | id_a__U72,id_nil => false | id_a__U72,id_o => false | id_a__U72,id_a__U52 => false | id_a__U72,id_U52 => false | id_a__U72,id_a__U23 => false | id_a__U72,id_U23 => false | id_a__U72,id_a__isPalListKind => false | id_a__U72,id_a__isNeList => false | id_a__U72,id_isNeList => false | id_a__U72,id_a__U71 => false | id_a__U72,id_U71 => false | id_a__U72,id_a__U41 => false | id_a__U72,id_U41 => false | id_a__U72,id_a__isPal => false | id_U72,id_a____ => true | id_U72,id_a => true | id_U72,id_a__U42 => true | id_U72,id_U42 => true | id_U72,id_a__U21 => true | id_U72,id_U21 => true | id_U72,id_a__U72 => true | id_U72,id_U72 => true | id_U72,id_a__U11 => false | id_U72,id_u => false | id_U72,id_a__U53 => false | id_U72,id_U53 => false | id_U72,id_a__U31 => false | id_U72,id_U31 => false | id_U72,id_isPalListKind => false | id_U72,id_mark => false | id_U72,id_i => false | id_U72,id_a__U51 => false | id_U72,id_U51 => false | id_U72,id_a__isList => false | id_U72,id_isList => false | id_U72,id_a__and => false | id_U72,id_a__U12 => false | id_U72,id_U12 => false | id_U72,id_a__U62 => false | id_U72,id_U62 => false | id_U72,id_a__isQid => false | id_U72,id_isQid => false | id_U72,id_isPal => false | id_U72,id___ => false | id_U72,id_e => false | id_U72,id_a__U43 => false | id_U72,id_U43 => false | id_U72,id_a__U22 => false | id_U72,id_U22 => false | id_U72,id_a__isNePal => false | id_U72,id_isNePal => false | id_U72,id_tt => false | id_U72,id_U11 => false | id_U72,id_a__U61 => false | id_U72,id_U61 => false | id_U72,id_a__U32 => false | id_U72,id_U32 => false | id_U72,id_and => false | id_U72,id_nil => false | id_U72,id_o => false | id_U72,id_a__U52 => false | id_U72,id_U52 => false | id_U72,id_a__U23 => false | id_U72,id_U23 => false | id_U72,id_a__isPalListKind => false | id_U72,id_a__isNeList => false | id_U72,id_isNeList => false | id_U72,id_a__U71 => false | id_U72,id_U71 => false | id_U72,id_a__U41 => false | id_U72,id_U41 => false | id_U72,id_a__isPal => false | id_a__U11,id_a____ => true | id_a__U11,id_a => true | id_a__U11,id_a__U42 => true | id_a__U11,id_U42 => true | id_a__U11,id_a__U21 => true | id_a__U11,id_U21 => true | id_a__U11,id_a__U72 => true | id_a__U11,id_U72 => true | id_a__U11,id_a__U11 => true | id_a__U11,id_u => false | id_a__U11,id_a__U53 => false | id_a__U11,id_U53 => false | id_a__U11,id_a__U31 => false | id_a__U11,id_U31 => false | id_a__U11,id_isPalListKind => false | id_a__U11,id_mark => false | id_a__U11,id_i => false | id_a__U11,id_a__U51 => false | id_a__U11,id_U51 => false | id_a__U11,id_a__isList => false | id_a__U11,id_isList => false | id_a__U11,id_a__and => false | id_a__U11,id_a__U12 => false | id_a__U11,id_U12 => false | id_a__U11,id_a__U62 => false | id_a__U11,id_U62 => false | id_a__U11,id_a__isQid => false | id_a__U11,id_isQid => false | id_a__U11,id_isPal => false | id_a__U11,id___ => false | id_a__U11,id_e => false | id_a__U11,id_a__U43 => false | id_a__U11,id_U43 => false | id_a__U11,id_a__U22 => false | id_a__U11,id_U22 => false | id_a__U11,id_a__isNePal => false | id_a__U11,id_isNePal => false | id_a__U11,id_tt => false | id_a__U11,id_U11 => false | id_a__U11,id_a__U61 => false | id_a__U11,id_U61 => false | id_a__U11,id_a__U32 => false | id_a__U11,id_U32 => false | id_a__U11,id_and => false | id_a__U11,id_nil => false | id_a__U11,id_o => false | id_a__U11,id_a__U52 => false | id_a__U11,id_U52 => false | id_a__U11,id_a__U23 => false | id_a__U11,id_U23 => false | id_a__U11,id_a__isPalListKind => false | id_a__U11,id_a__isNeList => false | id_a__U11,id_isNeList => false | id_a__U11,id_a__U71 => false | id_a__U11,id_U71 => false | id_a__U11,id_a__U41 => false | id_a__U11,id_U41 => false | id_a__U11,id_a__isPal => false | id_u,id_a____ => true | id_u,id_a => true | id_u,id_a__U42 => true | id_u,id_U42 => true | id_u,id_a__U21 => true | id_u,id_U21 => true | id_u,id_a__U72 => true | id_u,id_U72 => true | id_u,id_a__U11 => true | id_u,id_u => true | id_u,id_a__U53 => false | id_u,id_U53 => false | id_u,id_a__U31 => false | id_u,id_U31 => false | id_u,id_isPalListKind => false | id_u,id_mark => false | id_u,id_i => false | id_u,id_a__U51 => false | id_u,id_U51 => false | id_u,id_a__isList => false | id_u,id_isList => false | id_u,id_a__and => false | id_u,id_a__U12 => false | id_u,id_U12 => false | id_u,id_a__U62 => false | id_u,id_U62 => false | id_u,id_a__isQid => false | id_u,id_isQid => false | id_u,id_isPal => false | id_u,id___ => false | id_u,id_e => false | id_u,id_a__U43 => false | id_u,id_U43 => false | id_u,id_a__U22 => false | id_u,id_U22 => false | id_u,id_a__isNePal => false | id_u,id_isNePal => false | id_u,id_tt => false | id_u,id_U11 => false | id_u,id_a__U61 => false | id_u,id_U61 => false | id_u,id_a__U32 => false | id_u,id_U32 => false | id_u,id_and => false | id_u,id_nil => false | id_u,id_o => false | id_u,id_a__U52 => false | id_u,id_U52 => false | id_u,id_a__U23 => false | id_u,id_U23 => false | id_u,id_a__isPalListKind => false | id_u,id_a__isNeList => false | id_u,id_isNeList => false | id_u,id_a__U71 => false | id_u,id_U71 => false | id_u,id_a__U41 => false | id_u,id_U41 => false | id_u,id_a__isPal => false | id_a__U53,id_a____ => true | id_a__U53,id_a => true | id_a__U53,id_a__U42 => true | id_a__U53,id_U42 => true | id_a__U53,id_a__U21 => true | id_a__U53,id_U21 => true | id_a__U53,id_a__U72 => true | id_a__U53,id_U72 => true | id_a__U53,id_a__U11 => true | id_a__U53,id_u => true | id_a__U53,id_a__U53 => true | id_a__U53,id_U53 => false | id_a__U53,id_a__U31 => false | id_a__U53,id_U31 => false | id_a__U53,id_isPalListKind => false | id_a__U53,id_mark => false | id_a__U53,id_i => false | id_a__U53,id_a__U51 => false | id_a__U53,id_U51 => false | id_a__U53,id_a__isList => false | id_a__U53,id_isList => false | id_a__U53,id_a__and => false | id_a__U53,id_a__U12 => false | id_a__U53,id_U12 => false | id_a__U53,id_a__U62 => false | id_a__U53,id_U62 => false | id_a__U53,id_a__isQid => false | id_a__U53,id_isQid => false | id_a__U53,id_isPal => false | id_a__U53,id___ => false | id_a__U53,id_e => false | id_a__U53,id_a__U43 => false | id_a__U53,id_U43 => false | id_a__U53,id_a__U22 => false | id_a__U53,id_U22 => false | id_a__U53,id_a__isNePal => false | id_a__U53,id_isNePal => false | id_a__U53,id_tt => false | id_a__U53,id_U11 => false | id_a__U53,id_a__U61 => false | id_a__U53,id_U61 => false | id_a__U53,id_a__U32 => false | id_a__U53,id_U32 => false | id_a__U53,id_and => false | id_a__U53,id_nil => false | id_a__U53,id_o => false | id_a__U53,id_a__U52 => false | id_a__U53,id_U52 => false | id_a__U53,id_a__U23 => false | id_a__U53,id_U23 => false | id_a__U53,id_a__isPalListKind => false | id_a__U53,id_a__isNeList => false | id_a__U53,id_isNeList => false | id_a__U53,id_a__U71 => false | id_a__U53,id_U71 => false | id_a__U53,id_a__U41 => false | id_a__U53,id_U41 => false | id_a__U53,id_a__isPal => false | id_U53,id_a____ => true | id_U53,id_a => true | id_U53,id_a__U42 => true | id_U53,id_U42 => true | id_U53,id_a__U21 => true | id_U53,id_U21 => true | id_U53,id_a__U72 => true | id_U53,id_U72 => true | id_U53,id_a__U11 => true | id_U53,id_u => true | id_U53,id_a__U53 => true | id_U53,id_U53 => true | id_U53,id_a__U31 => false | id_U53,id_U31 => false | id_U53,id_isPalListKind => false | id_U53,id_mark => false | id_U53,id_i => false | id_U53,id_a__U51 => false | id_U53,id_U51 => false | id_U53,id_a__isList => false | id_U53,id_isList => false | id_U53,id_a__and => false | id_U53,id_a__U12 => false | id_U53,id_U12 => false | id_U53,id_a__U62 => false | id_U53,id_U62 => false | id_U53,id_a__isQid => false | id_U53,id_isQid => false | id_U53,id_isPal => false | id_U53,id___ => false | id_U53,id_e => false | id_U53,id_a__U43 => false | id_U53,id_U43 => false | id_U53,id_a__U22 => false | id_U53,id_U22 => false | id_U53,id_a__isNePal => false | id_U53,id_isNePal => false | id_U53,id_tt => false | id_U53,id_U11 => false | id_U53,id_a__U61 => false | id_U53,id_U61 => false | id_U53,id_a__U32 => false | id_U53,id_U32 => false | id_U53,id_and => false | id_U53,id_nil => false | id_U53,id_o => false | id_U53,id_a__U52 => false | id_U53,id_U52 => false | id_U53,id_a__U23 => false | id_U53,id_U23 => false | id_U53,id_a__isPalListKind => false | id_U53,id_a__isNeList => false | id_U53,id_isNeList => false | id_U53,id_a__U71 => false | id_U53,id_U71 => false | id_U53,id_a__U41 => false | id_U53,id_U41 => false | id_U53,id_a__isPal => false | id_a__U31,id_a____ => true | id_a__U31,id_a => true | id_a__U31,id_a__U42 => true | id_a__U31,id_U42 => true | id_a__U31,id_a__U21 => true | id_a__U31,id_U21 => true | id_a__U31,id_a__U72 => true | id_a__U31,id_U72 => true | id_a__U31,id_a__U11 => true | id_a__U31,id_u => true | id_a__U31,id_a__U53 => true | id_a__U31,id_U53 => true | id_a__U31,id_a__U31 => true | id_a__U31,id_U31 => false | id_a__U31,id_isPalListKind => false | id_a__U31,id_mark => false | id_a__U31,id_i => false | id_a__U31,id_a__U51 => false | id_a__U31,id_U51 => false | id_a__U31,id_a__isList => false | id_a__U31,id_isList => false | id_a__U31,id_a__and => false | id_a__U31,id_a__U12 => false | id_a__U31,id_U12 => false | id_a__U31,id_a__U62 => false | id_a__U31,id_U62 => false | id_a__U31,id_a__isQid => false | id_a__U31,id_isQid => false | id_a__U31,id_isPal => false | id_a__U31,id___ => false | id_a__U31,id_e => false | id_a__U31,id_a__U43 => false | id_a__U31,id_U43 => false | id_a__U31,id_a__U22 => false | id_a__U31,id_U22 => false | id_a__U31,id_a__isNePal => false | id_a__U31,id_isNePal => false | id_a__U31,id_tt => false | id_a__U31,id_U11 => false | id_a__U31,id_a__U61 => false | id_a__U31,id_U61 => false | id_a__U31,id_a__U32 => false | id_a__U31,id_U32 => false | id_a__U31,id_and => false | id_a__U31,id_nil => false | id_a__U31,id_o => false | id_a__U31,id_a__U52 => false | id_a__U31,id_U52 => false | id_a__U31,id_a__U23 => false | id_a__U31,id_U23 => false | id_a__U31,id_a__isPalListKind => false | id_a__U31,id_a__isNeList => false | id_a__U31,id_isNeList => false | id_a__U31,id_a__U71 => false | id_a__U31,id_U71 => false | id_a__U31,id_a__U41 => false | id_a__U31,id_U41 => false | id_a__U31,id_a__isPal => false | id_U31,id_a____ => true | id_U31,id_a => true | id_U31,id_a__U42 => true | id_U31,id_U42 => true | id_U31,id_a__U21 => true | id_U31,id_U21 => true | id_U31,id_a__U72 => true | id_U31,id_U72 => true | id_U31,id_a__U11 => true | id_U31,id_u => true | id_U31,id_a__U53 => true | id_U31,id_U53 => true | id_U31,id_a__U31 => true | id_U31,id_U31 => true | id_U31,id_isPalListKind => false | id_U31,id_mark => false | id_U31,id_i => false | id_U31,id_a__U51 => false | id_U31,id_U51 => false | id_U31,id_a__isList => false | id_U31,id_isList => false | id_U31,id_a__and => false | id_U31,id_a__U12 => false | id_U31,id_U12 => false | id_U31,id_a__U62 => false | id_U31,id_U62 => false | id_U31,id_a__isQid => false | id_U31,id_isQid => false | id_U31,id_isPal => false | id_U31,id___ => false | id_U31,id_e => false | id_U31,id_a__U43 => false | id_U31,id_U43 => false | id_U31,id_a__U22 => false | id_U31,id_U22 => false | id_U31,id_a__isNePal => false | id_U31,id_isNePal => false | id_U31,id_tt => false | id_U31,id_U11 => false | id_U31,id_a__U61 => false | id_U31,id_U61 => false | id_U31,id_a__U32 => false | id_U31,id_U32 => false | id_U31,id_and => false | id_U31,id_nil => false | id_U31,id_o => false | id_U31,id_a__U52 => false | id_U31,id_U52 => false | id_U31,id_a__U23 => false | id_U31,id_U23 => false | id_U31,id_a__isPalListKind => false | id_U31,id_a__isNeList => false | id_U31,id_isNeList => false | id_U31,id_a__U71 => false | id_U31,id_U71 => false | id_U31,id_a__U41 => false | id_U31,id_U41 => false | id_U31,id_a__isPal => false | id_isPalListKind,id_a____ => true | id_isPalListKind,id_a => true | id_isPalListKind,id_a__U42 => true | id_isPalListKind,id_U42 => true | id_isPalListKind,id_a__U21 => true | id_isPalListKind,id_U21 => true | id_isPalListKind,id_a__U72 => true | id_isPalListKind,id_U72 => true | id_isPalListKind,id_a__U11 => true | id_isPalListKind,id_u => true | id_isPalListKind,id_a__U53 => true | id_isPalListKind,id_U53 => true | id_isPalListKind,id_a__U31 => true | id_isPalListKind,id_U31 => true | id_isPalListKind,id_isPalListKind => true | id_isPalListKind,id_mark => false | id_isPalListKind,id_i => false | id_isPalListKind,id_a__U51 => false | id_isPalListKind,id_U51 => false | id_isPalListKind,id_a__isList => false | id_isPalListKind,id_isList => false | id_isPalListKind,id_a__and => false | id_isPalListKind,id_a__U12 => false | id_isPalListKind,id_U12 => false | id_isPalListKind,id_a__U62 => false | id_isPalListKind,id_U62 => false | id_isPalListKind,id_a__isQid => false | id_isPalListKind,id_isQid => false | id_isPalListKind,id_isPal => false | id_isPalListKind,id___ => false | id_isPalListKind,id_e => false | id_isPalListKind,id_a__U43 => false | id_isPalListKind,id_U43 => false | id_isPalListKind,id_a__U22 => false | id_isPalListKind,id_U22 => false | id_isPalListKind,id_a__isNePal => false | id_isPalListKind,id_isNePal => false | id_isPalListKind,id_tt => false | id_isPalListKind,id_U11 => false | id_isPalListKind,id_a__U61 => false | id_isPalListKind,id_U61 => false | id_isPalListKind,id_a__U32 => false | id_isPalListKind,id_U32 => false | id_isPalListKind,id_and => false | id_isPalListKind,id_nil => false | id_isPalListKind,id_o => false | id_isPalListKind,id_a__U52 => false | id_isPalListKind,id_U52 => false | id_isPalListKind,id_a__U23 => false | id_isPalListKind,id_U23 => false | id_isPalListKind,id_a__isPalListKind => false | id_isPalListKind,id_a__isNeList => false | id_isPalListKind,id_isNeList => false | id_isPalListKind,id_a__U71 => false | id_isPalListKind,id_U71 => false | id_isPalListKind,id_a__U41 => false | id_isPalListKind,id_U41 => false | id_isPalListKind,id_a__isPal => false | id_mark,id_a____ => true | id_mark,id_a => true | id_mark,id_a__U42 => true | id_mark,id_U42 => true | id_mark,id_a__U21 => true | id_mark,id_U21 => true | id_mark,id_a__U72 => true | id_mark,id_U72 => true | id_mark,id_a__U11 => true | id_mark,id_u => true | id_mark,id_a__U53 => true | id_mark,id_U53 => true | id_mark,id_a__U31 => true | id_mark,id_U31 => true | id_mark,id_isPalListKind => true | id_mark,id_mark => true | id_mark,id_i => false | id_mark,id_a__U51 => false | id_mark,id_U51 => false | id_mark,id_a__isList => false | id_mark,id_isList => false | id_mark,id_a__and => false | id_mark,id_a__U12 => false | id_mark,id_U12 => false | id_mark,id_a__U62 => false | id_mark,id_U62 => false | id_mark,id_a__isQid => false | id_mark,id_isQid => false | id_mark,id_isPal => false | id_mark,id___ => false | id_mark,id_e => false | id_mark,id_a__U43 => false | id_mark,id_U43 => false | id_mark,id_a__U22 => false | id_mark,id_U22 => false | id_mark,id_a__isNePal => false | id_mark,id_isNePal => false | id_mark,id_tt => false | id_mark,id_U11 => false | id_mark,id_a__U61 => false | id_mark,id_U61 => false | id_mark,id_a__U32 => false | id_mark,id_U32 => false | id_mark,id_and => false | id_mark,id_nil => false | id_mark,id_o => false | id_mark,id_a__U52 => false | id_mark,id_U52 => false | id_mark,id_a__U23 => false | id_mark,id_U23 => false | id_mark,id_a__isPalListKind => false | id_mark,id_a__isNeList => false | id_mark,id_isNeList => false | id_mark,id_a__U71 => false | id_mark,id_U71 => false | id_mark,id_a__U41 => false | id_mark,id_U41 => false | id_mark,id_a__isPal => false | id_i,id_a____ => true | id_i,id_a => true | id_i,id_a__U42 => true | id_i,id_U42 => true | id_i,id_a__U21 => true | id_i,id_U21 => true | id_i,id_a__U72 => true | id_i,id_U72 => true | id_i,id_a__U11 => true | id_i,id_u => true | id_i,id_a__U53 => true | id_i,id_U53 => true | id_i,id_a__U31 => true | id_i,id_U31 => true | id_i,id_isPalListKind => true | id_i,id_mark => true | id_i,id_i => true | id_i,id_a__U51 => false | id_i,id_U51 => false | id_i,id_a__isList => false | id_i,id_isList => false | id_i,id_a__and => false | id_i,id_a__U12 => false | id_i,id_U12 => false | id_i,id_a__U62 => false | id_i,id_U62 => false | id_i,id_a__isQid => false | id_i,id_isQid => false | id_i,id_isPal => false | id_i,id___ => false | id_i,id_e => false | id_i,id_a__U43 => false | id_i,id_U43 => false | id_i,id_a__U22 => false | id_i,id_U22 => false | id_i,id_a__isNePal => false | id_i,id_isNePal => false | id_i,id_tt => false | id_i,id_U11 => false | id_i,id_a__U61 => false | id_i,id_U61 => false | id_i,id_a__U32 => false | id_i,id_U32 => false | id_i,id_and => false | id_i,id_nil => false | id_i,id_o => false | id_i,id_a__U52 => false | id_i,id_U52 => false | id_i,id_a__U23 => false | id_i,id_U23 => false | id_i,id_a__isPalListKind => false | id_i,id_a__isNeList => false | id_i,id_isNeList => false | id_i,id_a__U71 => false | id_i,id_U71 => false | id_i,id_a__U41 => false | id_i,id_U41 => false | id_i,id_a__isPal => false | id_a__U51,id_a____ => true | id_a__U51,id_a => true | id_a__U51,id_a__U42 => true | id_a__U51,id_U42 => true | id_a__U51,id_a__U21 => true | id_a__U51,id_U21 => true | id_a__U51,id_a__U72 => true | id_a__U51,id_U72 => true | id_a__U51,id_a__U11 => true | id_a__U51,id_u => true | id_a__U51,id_a__U53 => true | id_a__U51,id_U53 => true | id_a__U51,id_a__U31 => true | id_a__U51,id_U31 => true | id_a__U51,id_isPalListKind => true | id_a__U51,id_mark => true | id_a__U51,id_i => true | id_a__U51,id_a__U51 => true | id_a__U51,id_U51 => false | id_a__U51,id_a__isList => false | id_a__U51,id_isList => false | id_a__U51,id_a__and => false | id_a__U51,id_a__U12 => false | id_a__U51,id_U12 => false | id_a__U51,id_a__U62 => false | id_a__U51,id_U62 => false | id_a__U51,id_a__isQid => false | id_a__U51,id_isQid => false | id_a__U51,id_isPal => false | id_a__U51,id___ => false | id_a__U51,id_e => false | id_a__U51,id_a__U43 => false | id_a__U51,id_U43 => false | id_a__U51,id_a__U22 => false | id_a__U51,id_U22 => false | id_a__U51,id_a__isNePal => false | id_a__U51,id_isNePal => false | id_a__U51,id_tt => false | id_a__U51,id_U11 => false | id_a__U51,id_a__U61 => false | id_a__U51,id_U61 => false | id_a__U51,id_a__U32 => false | id_a__U51,id_U32 => false | id_a__U51,id_and => false | id_a__U51,id_nil => false | id_a__U51,id_o => false | id_a__U51,id_a__U52 => false | id_a__U51,id_U52 => false | id_a__U51,id_a__U23 => false | id_a__U51,id_U23 => false | id_a__U51,id_a__isPalListKind => false | id_a__U51,id_a__isNeList => false | id_a__U51,id_isNeList => false | id_a__U51,id_a__U71 => false | id_a__U51,id_U71 => false | id_a__U51,id_a__U41 => false | id_a__U51,id_U41 => false | id_a__U51,id_a__isPal => false | id_U51,id_a____ => true | id_U51,id_a => true | id_U51,id_a__U42 => true | id_U51,id_U42 => true | id_U51,id_a__U21 => true | id_U51,id_U21 => true | id_U51,id_a__U72 => true | id_U51,id_U72 => true | id_U51,id_a__U11 => true | id_U51,id_u => true | id_U51,id_a__U53 => true | id_U51,id_U53 => true | id_U51,id_a__U31 => true | id_U51,id_U31 => true | id_U51,id_isPalListKind => true | id_U51,id_mark => true | id_U51,id_i => true | id_U51,id_a__U51 => true | id_U51,id_U51 => true | id_U51,id_a__isList => false | id_U51,id_isList => false | id_U51,id_a__and => false | id_U51,id_a__U12 => false | id_U51,id_U12 => false | id_U51,id_a__U62 => false | id_U51,id_U62 => false | id_U51,id_a__isQid => false | id_U51,id_isQid => false | id_U51,id_isPal => false | id_U51,id___ => false | id_U51,id_e => false | id_U51,id_a__U43 => false | id_U51,id_U43 => false | id_U51,id_a__U22 => false | id_U51,id_U22 => false | id_U51,id_a__isNePal => false | id_U51,id_isNePal => false | id_U51,id_tt => false | id_U51,id_U11 => false | id_U51,id_a__U61 => false | id_U51,id_U61 => false | id_U51,id_a__U32 => false | id_U51,id_U32 => false | id_U51,id_and => false | id_U51,id_nil => false | id_U51,id_o => false | id_U51,id_a__U52 => false | id_U51,id_U52 => false | id_U51,id_a__U23 => false | id_U51,id_U23 => false | id_U51,id_a__isPalListKind => false | id_U51,id_a__isNeList => false | id_U51,id_isNeList => false | id_U51,id_a__U71 => false | id_U51,id_U71 => false | id_U51,id_a__U41 => false | id_U51,id_U41 => false | id_U51,id_a__isPal => false | id_a__isList,id_a____ => true | id_a__isList,id_a => true | id_a__isList,id_a__U42 => true | id_a__isList,id_U42 => true | id_a__isList,id_a__U21 => true | id_a__isList,id_U21 => true | id_a__isList,id_a__U72 => true | id_a__isList,id_U72 => true | id_a__isList,id_a__U11 => true | id_a__isList,id_u => true | id_a__isList,id_a__U53 => true | id_a__isList,id_U53 => true | id_a__isList,id_a__U31 => true | id_a__isList,id_U31 => true | id_a__isList,id_isPalListKind => true | id_a__isList,id_mark => true | id_a__isList,id_i => true | id_a__isList,id_a__U51 => true | id_a__isList,id_U51 => true | id_a__isList,id_a__isList => true | id_a__isList,id_isList => false | id_a__isList,id_a__and => false | id_a__isList,id_a__U12 => false | id_a__isList,id_U12 => false | id_a__isList,id_a__U62 => false | id_a__isList,id_U62 => false | id_a__isList,id_a__isQid => false | id_a__isList,id_isQid => false | id_a__isList,id_isPal => false | id_a__isList,id___ => false | id_a__isList,id_e => false | id_a__isList,id_a__U43 => false | id_a__isList,id_U43 => false | id_a__isList,id_a__U22 => false | id_a__isList,id_U22 => false | id_a__isList,id_a__isNePal => false | id_a__isList,id_isNePal => false | id_a__isList,id_tt => false | id_a__isList,id_U11 => false | id_a__isList,id_a__U61 => false | id_a__isList,id_U61 => false | id_a__isList,id_a__U32 => false | id_a__isList,id_U32 => false | id_a__isList,id_and => false | id_a__isList,id_nil => false | id_a__isList,id_o => false | id_a__isList,id_a__U52 => false | id_a__isList,id_U52 => false | id_a__isList,id_a__U23 => false | id_a__isList,id_U23 => false | id_a__isList,id_a__isPalListKind => false | id_a__isList,id_a__isNeList => false | id_a__isList,id_isNeList => false | id_a__isList,id_a__U71 => false | id_a__isList,id_U71 => false | id_a__isList,id_a__U41 => false | id_a__isList,id_U41 => false | id_a__isList,id_a__isPal => false | id_isList,id_a____ => true | id_isList,id_a => true | id_isList,id_a__U42 => true | id_isList,id_U42 => true | id_isList,id_a__U21 => true | id_isList,id_U21 => true | id_isList,id_a__U72 => true | id_isList,id_U72 => true | id_isList,id_a__U11 => true | id_isList,id_u => true | id_isList,id_a__U53 => true | id_isList,id_U53 => true | id_isList,id_a__U31 => true | id_isList,id_U31 => true | id_isList,id_isPalListKind => true | id_isList,id_mark => true | id_isList,id_i => true | id_isList,id_a__U51 => true | id_isList,id_U51 => true | id_isList,id_a__isList => true | id_isList,id_isList => true | id_isList,id_a__and => false | id_isList,id_a__U12 => false | id_isList,id_U12 => false | id_isList,id_a__U62 => false | id_isList,id_U62 => false | id_isList,id_a__isQid => false | id_isList,id_isQid => false | id_isList,id_isPal => false | id_isList,id___ => false | id_isList,id_e => false | id_isList,id_a__U43 => false | id_isList,id_U43 => false | id_isList,id_a__U22 => false | id_isList,id_U22 => false | id_isList,id_a__isNePal => false | id_isList,id_isNePal => false | id_isList,id_tt => false | id_isList,id_U11 => false | id_isList,id_a__U61 => false | id_isList,id_U61 => false | id_isList,id_a__U32 => false | id_isList,id_U32 => false | id_isList,id_and => false | id_isList,id_nil => false | id_isList,id_o => false | id_isList,id_a__U52 => false | id_isList,id_U52 => false | id_isList,id_a__U23 => false | id_isList,id_U23 => false | id_isList,id_a__isPalListKind => false | id_isList,id_a__isNeList => false | id_isList,id_isNeList => false | id_isList,id_a__U71 => false | id_isList,id_U71 => false | id_isList,id_a__U41 => false | id_isList,id_U41 => false | id_isList,id_a__isPal => false | id_a__and,id_a____ => true | id_a__and,id_a => true | id_a__and,id_a__U42 => true | id_a__and,id_U42 => true | id_a__and,id_a__U21 => true | id_a__and,id_U21 => true | id_a__and,id_a__U72 => true | id_a__and,id_U72 => true | id_a__and,id_a__U11 => true | id_a__and,id_u => true | id_a__and,id_a__U53 => true | id_a__and,id_U53 => true | id_a__and,id_a__U31 => true | id_a__and,id_U31 => true | id_a__and,id_isPalListKind => true | id_a__and,id_mark => true | id_a__and,id_i => true | id_a__and,id_a__U51 => true | id_a__and,id_U51 => true | id_a__and,id_a__isList => true | id_a__and,id_isList => true | id_a__and,id_a__and => true | id_a__and,id_a__U12 => false | id_a__and,id_U12 => false | id_a__and,id_a__U62 => false | id_a__and,id_U62 => false | id_a__and,id_a__isQid => false | id_a__and,id_isQid => false | id_a__and,id_isPal => false | id_a__and,id___ => false | id_a__and,id_e => false | id_a__and,id_a__U43 => false | id_a__and,id_U43 => false | id_a__and,id_a__U22 => false | id_a__and,id_U22 => false | id_a__and,id_a__isNePal => false | id_a__and,id_isNePal => false | id_a__and,id_tt => false | id_a__and,id_U11 => false | id_a__and,id_a__U61 => false | id_a__and,id_U61 => false | id_a__and,id_a__U32 => false | id_a__and,id_U32 => false | id_a__and,id_and => false | id_a__and,id_nil => false | id_a__and,id_o => false | id_a__and,id_a__U52 => false | id_a__and,id_U52 => false | id_a__and,id_a__U23 => false | id_a__and,id_U23 => false | id_a__and,id_a__isPalListKind => false | id_a__and,id_a__isNeList => false | id_a__and,id_isNeList => false | id_a__and,id_a__U71 => false | id_a__and,id_U71 => false | id_a__and,id_a__U41 => false | id_a__and,id_U41 => false | id_a__and,id_a__isPal => false | id_a__U12,id_a____ => true | id_a__U12,id_a => true | id_a__U12,id_a__U42 => true | id_a__U12,id_U42 => true | id_a__U12,id_a__U21 => true | id_a__U12,id_U21 => true | id_a__U12,id_a__U72 => true | id_a__U12,id_U72 => true | id_a__U12,id_a__U11 => true | id_a__U12,id_u => true | id_a__U12,id_a__U53 => true | id_a__U12,id_U53 => true | id_a__U12,id_a__U31 => true | id_a__U12,id_U31 => true | id_a__U12,id_isPalListKind => true | id_a__U12,id_mark => true | id_a__U12,id_i => true | id_a__U12,id_a__U51 => true | id_a__U12,id_U51 => true | id_a__U12,id_a__isList => true | id_a__U12,id_isList => true | id_a__U12,id_a__and => true | id_a__U12,id_a__U12 => true | id_a__U12,id_U12 => false | id_a__U12,id_a__U62 => false | id_a__U12,id_U62 => false | id_a__U12,id_a__isQid => false | id_a__U12,id_isQid => false | id_a__U12,id_isPal => false | id_a__U12,id___ => false | id_a__U12,id_e => false | id_a__U12,id_a__U43 => false | id_a__U12,id_U43 => false | id_a__U12,id_a__U22 => false | id_a__U12,id_U22 => false | id_a__U12,id_a__isNePal => false | id_a__U12,id_isNePal => false | id_a__U12,id_tt => false | id_a__U12,id_U11 => false | id_a__U12,id_a__U61 => false | id_a__U12,id_U61 => false | id_a__U12,id_a__U32 => false | id_a__U12,id_U32 => false | id_a__U12,id_and => false | id_a__U12,id_nil => false | id_a__U12,id_o => false | id_a__U12,id_a__U52 => false | id_a__U12,id_U52 => false | id_a__U12,id_a__U23 => false | id_a__U12,id_U23 => false | id_a__U12,id_a__isPalListKind => false | id_a__U12,id_a__isNeList => false | id_a__U12,id_isNeList => false | id_a__U12,id_a__U71 => false | id_a__U12,id_U71 => false | id_a__U12,id_a__U41 => false | id_a__U12,id_U41 => false | id_a__U12,id_a__isPal => false | id_U12,id_a____ => true | id_U12,id_a => true | id_U12,id_a__U42 => true | id_U12,id_U42 => true | id_U12,id_a__U21 => true | id_U12,id_U21 => true | id_U12,id_a__U72 => true | id_U12,id_U72 => true | id_U12,id_a__U11 => true | id_U12,id_u => true | id_U12,id_a__U53 => true | id_U12,id_U53 => true | id_U12,id_a__U31 => true | id_U12,id_U31 => true | id_U12,id_isPalListKind => true | id_U12,id_mark => true | id_U12,id_i => true | id_U12,id_a__U51 => true | id_U12,id_U51 => true | id_U12,id_a__isList => true | id_U12,id_isList => true | id_U12,id_a__and => true | id_U12,id_a__U12 => true | id_U12,id_U12 => true | id_U12,id_a__U62 => false | id_U12,id_U62 => false | id_U12,id_a__isQid => false | id_U12,id_isQid => false | id_U12,id_isPal => false | id_U12,id___ => false | id_U12,id_e => false | id_U12,id_a__U43 => false | id_U12,id_U43 => false | id_U12,id_a__U22 => false | id_U12,id_U22 => false | id_U12,id_a__isNePal => false | id_U12,id_isNePal => false | id_U12,id_tt => false | id_U12,id_U11 => false | id_U12,id_a__U61 => false | id_U12,id_U61 => false | id_U12,id_a__U32 => false | id_U12,id_U32 => false | id_U12,id_and => false | id_U12,id_nil => false | id_U12,id_o => false | id_U12,id_a__U52 => false | id_U12,id_U52 => false | id_U12,id_a__U23 => false | id_U12,id_U23 => false | id_U12,id_a__isPalListKind => false | id_U12,id_a__isNeList => false | id_U12,id_isNeList => false | id_U12,id_a__U71 => false | id_U12,id_U71 => false | id_U12,id_a__U41 => false | id_U12,id_U41 => false | id_U12,id_a__isPal => false | id_a__U62,id_a____ => true | id_a__U62,id_a => true | id_a__U62,id_a__U42 => true | id_a__U62,id_U42 => true | id_a__U62,id_a__U21 => true | id_a__U62,id_U21 => true | id_a__U62,id_a__U72 => true | id_a__U62,id_U72 => true | id_a__U62,id_a__U11 => true | id_a__U62,id_u => true | id_a__U62,id_a__U53 => true | id_a__U62,id_U53 => true | id_a__U62,id_a__U31 => true | id_a__U62,id_U31 => true | id_a__U62,id_isPalListKind => true | id_a__U62,id_mark => true | id_a__U62,id_i => true | id_a__U62,id_a__U51 => true | id_a__U62,id_U51 => true | id_a__U62,id_a__isList => true | id_a__U62,id_isList => true | id_a__U62,id_a__and => true | id_a__U62,id_a__U12 => true | id_a__U62,id_U12 => true | id_a__U62,id_a__U62 => true | id_a__U62,id_U62 => false | id_a__U62,id_a__isQid => false | id_a__U62,id_isQid => false | id_a__U62,id_isPal => false | id_a__U62,id___ => false | id_a__U62,id_e => false | id_a__U62,id_a__U43 => false | id_a__U62,id_U43 => false | id_a__U62,id_a__U22 => false | id_a__U62,id_U22 => false | id_a__U62,id_a__isNePal => false | id_a__U62,id_isNePal => false | id_a__U62,id_tt => false | id_a__U62,id_U11 => false | id_a__U62,id_a__U61 => false | id_a__U62,id_U61 => false | id_a__U62,id_a__U32 => false | id_a__U62,id_U32 => false | id_a__U62,id_and => false | id_a__U62,id_nil => false | id_a__U62,id_o => false | id_a__U62,id_a__U52 => false | id_a__U62,id_U52 => false | id_a__U62,id_a__U23 => false | id_a__U62,id_U23 => false | id_a__U62,id_a__isPalListKind => false | id_a__U62,id_a__isNeList => false | id_a__U62,id_isNeList => false | id_a__U62,id_a__U71 => false | id_a__U62,id_U71 => false | id_a__U62,id_a__U41 => false | id_a__U62,id_U41 => false | id_a__U62,id_a__isPal => false | id_U62,id_a____ => true | id_U62,id_a => true | id_U62,id_a__U42 => true | id_U62,id_U42 => true | id_U62,id_a__U21 => true | id_U62,id_U21 => true | id_U62,id_a__U72 => true | id_U62,id_U72 => true | id_U62,id_a__U11 => true | id_U62,id_u => true | id_U62,id_a__U53 => true | id_U62,id_U53 => true | id_U62,id_a__U31 => true | id_U62,id_U31 => true | id_U62,id_isPalListKind => true | id_U62,id_mark => true | id_U62,id_i => true | id_U62,id_a__U51 => true | id_U62,id_U51 => true | id_U62,id_a__isList => true | id_U62,id_isList => true | id_U62,id_a__and => true | id_U62,id_a__U12 => true | id_U62,id_U12 => true | id_U62,id_a__U62 => true | id_U62,id_U62 => true | id_U62,id_a__isQid => false | id_U62,id_isQid => false | id_U62,id_isPal => false | id_U62,id___ => false | id_U62,id_e => false | id_U62,id_a__U43 => false | id_U62,id_U43 => false | id_U62,id_a__U22 => false | id_U62,id_U22 => false | id_U62,id_a__isNePal => false | id_U62,id_isNePal => false | id_U62,id_tt => false | id_U62,id_U11 => false | id_U62,id_a__U61 => false | id_U62,id_U61 => false | id_U62,id_a__U32 => false | id_U62,id_U32 => false | id_U62,id_and => false | id_U62,id_nil => false | id_U62,id_o => false | id_U62,id_a__U52 => false | id_U62,id_U52 => false | id_U62,id_a__U23 => false | id_U62,id_U23 => false | id_U62,id_a__isPalListKind => false | id_U62,id_a__isNeList => false | id_U62,id_isNeList => false | id_U62,id_a__U71 => false | id_U62,id_U71 => false | id_U62,id_a__U41 => false | id_U62,id_U41 => false | id_U62,id_a__isPal => false | id_a__isQid,id_a____ => true | id_a__isQid,id_a => true | id_a__isQid,id_a__U42 => true | id_a__isQid,id_U42 => true | id_a__isQid,id_a__U21 => true | id_a__isQid,id_U21 => true | id_a__isQid,id_a__U72 => true | id_a__isQid,id_U72 => true | id_a__isQid,id_a__U11 => true | id_a__isQid,id_u => true | id_a__isQid,id_a__U53 => true | id_a__isQid,id_U53 => true | id_a__isQid,id_a__U31 => true | id_a__isQid,id_U31 => true | id_a__isQid,id_isPalListKind => true | id_a__isQid,id_mark => true | id_a__isQid,id_i => true | id_a__isQid,id_a__U51 => true | id_a__isQid,id_U51 => true | id_a__isQid,id_a__isList => true | id_a__isQid,id_isList => true | id_a__isQid,id_a__and => true | id_a__isQid,id_a__U12 => true | id_a__isQid,id_U12 => true | id_a__isQid,id_a__U62 => true | id_a__isQid,id_U62 => true | id_a__isQid,id_a__isQid => true | id_a__isQid,id_isQid => false | id_a__isQid,id_isPal => false | id_a__isQid,id___ => false | id_a__isQid,id_e => false | id_a__isQid,id_a__U43 => false | id_a__isQid,id_U43 => false | id_a__isQid,id_a__U22 => false | id_a__isQid,id_U22 => false | id_a__isQid,id_a__isNePal => false | id_a__isQid,id_isNePal => false | id_a__isQid,id_tt => false | id_a__isQid,id_U11 => false | id_a__isQid,id_a__U61 => false | id_a__isQid,id_U61 => false | id_a__isQid,id_a__U32 => false | id_a__isQid,id_U32 => false | id_a__isQid,id_and => false | id_a__isQid,id_nil => false | id_a__isQid,id_o => false | id_a__isQid,id_a__U52 => false | id_a__isQid,id_U52 => false | id_a__isQid,id_a__U23 => false | id_a__isQid,id_U23 => false | id_a__isQid,id_a__isPalListKind => false | id_a__isQid,id_a__isNeList => false | id_a__isQid,id_isNeList => false | id_a__isQid,id_a__U71 => false | id_a__isQid,id_U71 => false | id_a__isQid,id_a__U41 => false | id_a__isQid,id_U41 => false | id_a__isQid,id_a__isPal => false | id_isQid,id_a____ => true | id_isQid,id_a => true | id_isQid,id_a__U42 => true | id_isQid,id_U42 => true | id_isQid,id_a__U21 => true | id_isQid,id_U21 => true | id_isQid,id_a__U72 => true | id_isQid,id_U72 => true | id_isQid,id_a__U11 => true | id_isQid,id_u => true | id_isQid,id_a__U53 => true | id_isQid,id_U53 => true | id_isQid,id_a__U31 => true | id_isQid,id_U31 => true | id_isQid,id_isPalListKind => true | id_isQid,id_mark => true | id_isQid,id_i => true | id_isQid,id_a__U51 => true | id_isQid,id_U51 => true | id_isQid,id_a__isList => true | id_isQid,id_isList => true | id_isQid,id_a__and => true | id_isQid,id_a__U12 => true | id_isQid,id_U12 => true | id_isQid,id_a__U62 => true | id_isQid,id_U62 => true | id_isQid,id_a__isQid => true | id_isQid,id_isQid => true | id_isQid,id_isPal => false | id_isQid,id___ => false | id_isQid,id_e => false | id_isQid,id_a__U43 => false | id_isQid,id_U43 => false | id_isQid,id_a__U22 => false | id_isQid,id_U22 => false | id_isQid,id_a__isNePal => false | id_isQid,id_isNePal => false | id_isQid,id_tt => false | id_isQid,id_U11 => false | id_isQid,id_a__U61 => false | id_isQid,id_U61 => false | id_isQid,id_a__U32 => false | id_isQid,id_U32 => false | id_isQid,id_and => false | id_isQid,id_nil => false | id_isQid,id_o => false | id_isQid,id_a__U52 => false | id_isQid,id_U52 => false | id_isQid,id_a__U23 => false | id_isQid,id_U23 => false | id_isQid,id_a__isPalListKind => false | id_isQid,id_a__isNeList => false | id_isQid,id_isNeList => false | id_isQid,id_a__U71 => false | id_isQid,id_U71 => false | id_isQid,id_a__U41 => false | id_isQid,id_U41 => false | id_isQid,id_a__isPal => false | id_isPal,id_a____ => true | id_isPal,id_a => true | id_isPal,id_a__U42 => true | id_isPal,id_U42 => true | id_isPal,id_a__U21 => true | id_isPal,id_U21 => true | id_isPal,id_a__U72 => true | id_isPal,id_U72 => true | id_isPal,id_a__U11 => true | id_isPal,id_u => true | id_isPal,id_a__U53 => true | id_isPal,id_U53 => true | id_isPal,id_a__U31 => true | id_isPal,id_U31 => true | id_isPal,id_isPalListKind => true | id_isPal,id_mark => true | id_isPal,id_i => true | id_isPal,id_a__U51 => true | id_isPal,id_U51 => true | id_isPal,id_a__isList => true | id_isPal,id_isList => true | id_isPal,id_a__and => true | id_isPal,id_a__U12 => true | id_isPal,id_U12 => true | id_isPal,id_a__U62 => true | id_isPal,id_U62 => true | id_isPal,id_a__isQid => true | id_isPal,id_isQid => true | id_isPal,id_isPal => true | id_isPal,id___ => false | id_isPal,id_e => false | id_isPal,id_a__U43 => false | id_isPal,id_U43 => false | id_isPal,id_a__U22 => false | id_isPal,id_U22 => false | id_isPal,id_a__isNePal => false | id_isPal,id_isNePal => false | id_isPal,id_tt => false | id_isPal,id_U11 => false | id_isPal,id_a__U61 => false | id_isPal,id_U61 => false | id_isPal,id_a__U32 => false | id_isPal,id_U32 => false | id_isPal,id_and => false | id_isPal,id_nil => false | id_isPal,id_o => false | id_isPal,id_a__U52 => false | id_isPal,id_U52 => false | id_isPal,id_a__U23 => false | id_isPal,id_U23 => false | id_isPal,id_a__isPalListKind => false | id_isPal,id_a__isNeList => false | id_isPal,id_isNeList => false | id_isPal,id_a__U71 => false | id_isPal,id_U71 => false | id_isPal,id_a__U41 => false | id_isPal,id_U41 => false | id_isPal,id_a__isPal => false | id___,id_a____ => true | id___,id_a => true | id___,id_a__U42 => true | id___,id_U42 => true | id___,id_a__U21 => true | id___,id_U21 => true | id___,id_a__U72 => true | id___,id_U72 => true | id___,id_a__U11 => true | id___,id_u => true | id___,id_a__U53 => true | id___,id_U53 => true | id___,id_a__U31 => true | id___,id_U31 => true | id___,id_isPalListKind => true | id___,id_mark => true | id___,id_i => true | id___,id_a__U51 => true | id___,id_U51 => true | id___,id_a__isList => true | id___,id_isList => true | id___,id_a__and => true | id___,id_a__U12 => true | id___,id_U12 => true | id___,id_a__U62 => true | id___,id_U62 => true | id___,id_a__isQid => true | id___,id_isQid => true | id___,id_isPal => true | id___,id___ => true | id___,id_e => false | id___,id_a__U43 => false | id___,id_U43 => false | id___,id_a__U22 => false | id___,id_U22 => false | id___,id_a__isNePal => false | id___,id_isNePal => false | id___,id_tt => false | id___,id_U11 => false | id___,id_a__U61 => false | id___,id_U61 => false | id___,id_a__U32 => false | id___,id_U32 => false | id___,id_and => false | id___,id_nil => false | id___,id_o => false | id___,id_a__U52 => false | id___,id_U52 => false | id___,id_a__U23 => false | id___,id_U23 => false | id___,id_a__isPalListKind => false | id___,id_a__isNeList => false | id___,id_isNeList => false | id___,id_a__U71 => false | id___,id_U71 => false | id___,id_a__U41 => false | id___,id_U41 => false | id___,id_a__isPal => false | id_e,id_a____ => true | id_e,id_a => true | id_e,id_a__U42 => true | id_e,id_U42 => true | id_e,id_a__U21 => true | id_e,id_U21 => true | id_e,id_a__U72 => true | id_e,id_U72 => true | id_e,id_a__U11 => true | id_e,id_u => true | id_e,id_a__U53 => true | id_e,id_U53 => true | id_e,id_a__U31 => true | id_e,id_U31 => true | id_e,id_isPalListKind => true | id_e,id_mark => true | id_e,id_i => true | id_e,id_a__U51 => true | id_e,id_U51 => true | id_e,id_a__isList => true | id_e,id_isList => true | id_e,id_a__and => true | id_e,id_a__U12 => true | id_e,id_U12 => true | id_e,id_a__U62 => true | id_e,id_U62 => true | id_e,id_a__isQid => true | id_e,id_isQid => true | id_e,id_isPal => true | id_e,id___ => true | id_e,id_e => true | id_e,id_a__U43 => false | id_e,id_U43 => false | id_e,id_a__U22 => false | id_e,id_U22 => false | id_e,id_a__isNePal => false | id_e,id_isNePal => false | id_e,id_tt => false | id_e,id_U11 => false | id_e,id_a__U61 => false | id_e,id_U61 => false | id_e,id_a__U32 => false | id_e,id_U32 => false | id_e,id_and => false | id_e,id_nil => false | id_e,id_o => false | id_e,id_a__U52 => false | id_e,id_U52 => false | id_e,id_a__U23 => false | id_e,id_U23 => false | id_e,id_a__isPalListKind => false | id_e,id_a__isNeList => false | id_e,id_isNeList => false | id_e,id_a__U71 => false | id_e,id_U71 => false | id_e,id_a__U41 => false | id_e,id_U41 => false | id_e,id_a__isPal => false | id_a__U43,id_a____ => true | id_a__U43,id_a => true | id_a__U43,id_a__U42 => true | id_a__U43,id_U42 => true | id_a__U43,id_a__U21 => true | id_a__U43,id_U21 => true | id_a__U43,id_a__U72 => true | id_a__U43,id_U72 => true | id_a__U43,id_a__U11 => true | id_a__U43,id_u => true | id_a__U43,id_a__U53 => true | id_a__U43,id_U53 => true | id_a__U43,id_a__U31 => true | id_a__U43,id_U31 => true | id_a__U43,id_isPalListKind => true | id_a__U43,id_mark => true | id_a__U43,id_i => true | id_a__U43,id_a__U51 => true | id_a__U43,id_U51 => true | id_a__U43,id_a__isList => true | id_a__U43,id_isList => true | id_a__U43,id_a__and => true | id_a__U43,id_a__U12 => true | id_a__U43,id_U12 => true | id_a__U43,id_a__U62 => true | id_a__U43,id_U62 => true | id_a__U43,id_a__isQid => true | id_a__U43,id_isQid => true | id_a__U43,id_isPal => true | id_a__U43,id___ => true | id_a__U43,id_e => true | id_a__U43,id_a__U43 => true | id_a__U43,id_U43 => false | id_a__U43,id_a__U22 => false | id_a__U43,id_U22 => false | id_a__U43,id_a__isNePal => false | id_a__U43,id_isNePal => false | id_a__U43,id_tt => false | id_a__U43,id_U11 => false | id_a__U43,id_a__U61 => false | id_a__U43,id_U61 => false | id_a__U43,id_a__U32 => false | id_a__U43,id_U32 => false | id_a__U43,id_and => false | id_a__U43,id_nil => false | id_a__U43,id_o => false | id_a__U43,id_a__U52 => false | id_a__U43,id_U52 => false | id_a__U43,id_a__U23 => false | id_a__U43,id_U23 => false | id_a__U43,id_a__isPalListKind => false | id_a__U43,id_a__isNeList => false | id_a__U43,id_isNeList => false | id_a__U43,id_a__U71 => false | id_a__U43,id_U71 => false | id_a__U43,id_a__U41 => false | id_a__U43,id_U41 => false | id_a__U43,id_a__isPal => false | id_U43,id_a____ => true | id_U43,id_a => true | id_U43,id_a__U42 => true | id_U43,id_U42 => true | id_U43,id_a__U21 => true | id_U43,id_U21 => true | id_U43,id_a__U72 => true | id_U43,id_U72 => true | id_U43,id_a__U11 => true | id_U43,id_u => true | id_U43,id_a__U53 => true | id_U43,id_U53 => true | id_U43,id_a__U31 => true | id_U43,id_U31 => true | id_U43,id_isPalListKind => true | id_U43,id_mark => true | id_U43,id_i => true | id_U43,id_a__U51 => true | id_U43,id_U51 => true | id_U43,id_a__isList => true | id_U43,id_isList => true | id_U43,id_a__and => true | id_U43,id_a__U12 => true | id_U43,id_U12 => true | id_U43,id_a__U62 => true | id_U43,id_U62 => true | id_U43,id_a__isQid => true | id_U43,id_isQid => true | id_U43,id_isPal => true | id_U43,id___ => true | id_U43,id_e => true | id_U43,id_a__U43 => true | id_U43,id_U43 => true | id_U43,id_a__U22 => false | id_U43,id_U22 => false | id_U43,id_a__isNePal => false | id_U43,id_isNePal => false | id_U43,id_tt => false | id_U43,id_U11 => false | id_U43,id_a__U61 => false | id_U43,id_U61 => false | id_U43,id_a__U32 => false | id_U43,id_U32 => false | id_U43,id_and => false | id_U43,id_nil => false | id_U43,id_o => false | id_U43,id_a__U52 => false | id_U43,id_U52 => false | id_U43,id_a__U23 => false | id_U43,id_U23 => false | id_U43,id_a__isPalListKind => false | id_U43,id_a__isNeList => false | id_U43,id_isNeList => false | id_U43,id_a__U71 => false | id_U43,id_U71 => false | id_U43,id_a__U41 => false | id_U43,id_U41 => false | id_U43,id_a__isPal => false | id_a__U22,id_a____ => true | id_a__U22,id_a => true | id_a__U22,id_a__U42 => true | id_a__U22,id_U42 => true | id_a__U22,id_a__U21 => true | id_a__U22,id_U21 => true | id_a__U22,id_a__U72 => true | id_a__U22,id_U72 => true | id_a__U22,id_a__U11 => true | id_a__U22,id_u => true | id_a__U22,id_a__U53 => true | id_a__U22,id_U53 => true | id_a__U22,id_a__U31 => true | id_a__U22,id_U31 => true | id_a__U22,id_isPalListKind => true | id_a__U22,id_mark => true | id_a__U22,id_i => true | id_a__U22,id_a__U51 => true | id_a__U22,id_U51 => true | id_a__U22,id_a__isList => true | id_a__U22,id_isList => true | id_a__U22,id_a__and => true | id_a__U22,id_a__U12 => true | id_a__U22,id_U12 => true | id_a__U22,id_a__U62 => true | id_a__U22,id_U62 => true | id_a__U22,id_a__isQid => true | id_a__U22,id_isQid => true | id_a__U22,id_isPal => true | id_a__U22,id___ => true | id_a__U22,id_e => true | id_a__U22,id_a__U43 => true | id_a__U22,id_U43 => true | id_a__U22,id_a__U22 => true | id_a__U22,id_U22 => false | id_a__U22,id_a__isNePal => false | id_a__U22,id_isNePal => false | id_a__U22,id_tt => false | id_a__U22,id_U11 => false | id_a__U22,id_a__U61 => false | id_a__U22,id_U61 => false | id_a__U22,id_a__U32 => false | id_a__U22,id_U32 => false | id_a__U22,id_and => false | id_a__U22,id_nil => false | id_a__U22,id_o => false | id_a__U22,id_a__U52 => false | id_a__U22,id_U52 => false | id_a__U22,id_a__U23 => false | id_a__U22,id_U23 => false | id_a__U22,id_a__isPalListKind => false | id_a__U22,id_a__isNeList => false | id_a__U22,id_isNeList => false | id_a__U22,id_a__U71 => false | id_a__U22,id_U71 => false | id_a__U22,id_a__U41 => false | id_a__U22,id_U41 => false | id_a__U22,id_a__isPal => false | id_U22,id_a____ => true | id_U22,id_a => true | id_U22,id_a__U42 => true | id_U22,id_U42 => true | id_U22,id_a__U21 => true | id_U22,id_U21 => true | id_U22,id_a__U72 => true | id_U22,id_U72 => true | id_U22,id_a__U11 => true | id_U22,id_u => true | id_U22,id_a__U53 => true | id_U22,id_U53 => true | id_U22,id_a__U31 => true | id_U22,id_U31 => true | id_U22,id_isPalListKind => true | id_U22,id_mark => true | id_U22,id_i => true | id_U22,id_a__U51 => true | id_U22,id_U51 => true | id_U22,id_a__isList => true | id_U22,id_isList => true | id_U22,id_a__and => true | id_U22,id_a__U12 => true | id_U22,id_U12 => true | id_U22,id_a__U62 => true | id_U22,id_U62 => true | id_U22,id_a__isQid => true | id_U22,id_isQid => true | id_U22,id_isPal => true | id_U22,id___ => true | id_U22,id_e => true | id_U22,id_a__U43 => true | id_U22,id_U43 => true | id_U22,id_a__U22 => true | id_U22,id_U22 => true | id_U22,id_a__isNePal => false | id_U22,id_isNePal => false | id_U22,id_tt => false | id_U22,id_U11 => false | id_U22,id_a__U61 => false | id_U22,id_U61 => false | id_U22,id_a__U32 => false | id_U22,id_U32 => false | id_U22,id_and => false | id_U22,id_nil => false | id_U22,id_o => false | id_U22,id_a__U52 => false | id_U22,id_U52 => false | id_U22,id_a__U23 => false | id_U22,id_U23 => false | id_U22,id_a__isPalListKind => false | id_U22,id_a__isNeList => false | id_U22,id_isNeList => false | id_U22,id_a__U71 => false | id_U22,id_U71 => false | id_U22,id_a__U41 => false | id_U22,id_U41 => false | id_U22,id_a__isPal => false | id_a__isNePal,id_a____ => true | id_a__isNePal,id_a => true | id_a__isNePal,id_a__U42 => true | id_a__isNePal,id_U42 => true | id_a__isNePal,id_a__U21 => true | id_a__isNePal,id_U21 => true | id_a__isNePal,id_a__U72 => true | id_a__isNePal,id_U72 => true | id_a__isNePal,id_a__U11 => true | id_a__isNePal,id_u => true | id_a__isNePal,id_a__U53 => true | id_a__isNePal,id_U53 => true | id_a__isNePal,id_a__U31 => true | id_a__isNePal,id_U31 => true | id_a__isNePal,id_isPalListKind => true | id_a__isNePal,id_mark => true | id_a__isNePal,id_i => true | id_a__isNePal,id_a__U51 => true | id_a__isNePal,id_U51 => true | id_a__isNePal,id_a__isList => true | id_a__isNePal,id_isList => true | id_a__isNePal,id_a__and => true | id_a__isNePal,id_a__U12 => true | id_a__isNePal,id_U12 => true | id_a__isNePal,id_a__U62 => true | id_a__isNePal,id_U62 => true | id_a__isNePal,id_a__isQid => true | id_a__isNePal,id_isQid => true | id_a__isNePal,id_isPal => true | id_a__isNePal,id___ => true | id_a__isNePal,id_e => true | id_a__isNePal,id_a__U43 => true | id_a__isNePal,id_U43 => true | id_a__isNePal,id_a__U22 => true | id_a__isNePal,id_U22 => true | id_a__isNePal,id_a__isNePal => true | id_a__isNePal,id_isNePal => false | id_a__isNePal,id_tt => false | id_a__isNePal,id_U11 => false | id_a__isNePal,id_a__U61 => false | id_a__isNePal,id_U61 => false | id_a__isNePal,id_a__U32 => false | id_a__isNePal,id_U32 => false | id_a__isNePal,id_and => false | id_a__isNePal,id_nil => false | id_a__isNePal,id_o => false | id_a__isNePal,id_a__U52 => false | id_a__isNePal,id_U52 => false | id_a__isNePal,id_a__U23 => false | id_a__isNePal,id_U23 => false | id_a__isNePal,id_a__isPalListKind => false | id_a__isNePal,id_a__isNeList => false | id_a__isNePal,id_isNeList => false | id_a__isNePal,id_a__U71 => false | id_a__isNePal,id_U71 => false | id_a__isNePal,id_a__U41 => false | id_a__isNePal,id_U41 => false | id_a__isNePal,id_a__isPal => false | id_isNePal,id_a____ => true | id_isNePal,id_a => true | id_isNePal,id_a__U42 => true | id_isNePal,id_U42 => true | id_isNePal,id_a__U21 => true | id_isNePal,id_U21 => true | id_isNePal,id_a__U72 => true | id_isNePal,id_U72 => true | id_isNePal,id_a__U11 => true | id_isNePal,id_u => true | id_isNePal,id_a__U53 => true | id_isNePal,id_U53 => true | id_isNePal,id_a__U31 => true | id_isNePal,id_U31 => true | id_isNePal,id_isPalListKind => true | id_isNePal,id_mark => true | id_isNePal,id_i => true | id_isNePal,id_a__U51 => true | id_isNePal,id_U51 => true | id_isNePal,id_a__isList => true | id_isNePal,id_isList => true | id_isNePal,id_a__and => true | id_isNePal,id_a__U12 => true | id_isNePal,id_U12 => true | id_isNePal,id_a__U62 => true | id_isNePal,id_U62 => true | id_isNePal,id_a__isQid => true | id_isNePal,id_isQid => true | id_isNePal,id_isPal => true | id_isNePal,id___ => true | id_isNePal,id_e => true | id_isNePal,id_a__U43 => true | id_isNePal,id_U43 => true | id_isNePal,id_a__U22 => true | id_isNePal,id_U22 => true | id_isNePal,id_a__isNePal => true | id_isNePal,id_isNePal => true | id_isNePal,id_tt => false | id_isNePal,id_U11 => false | id_isNePal,id_a__U61 => false | id_isNePal,id_U61 => false | id_isNePal,id_a__U32 => false | id_isNePal,id_U32 => false | id_isNePal,id_and => false | id_isNePal,id_nil => false | id_isNePal,id_o => false | id_isNePal,id_a__U52 => false | id_isNePal,id_U52 => false | id_isNePal,id_a__U23 => false | id_isNePal,id_U23 => false | id_isNePal,id_a__isPalListKind => false | id_isNePal,id_a__isNeList => false | id_isNePal,id_isNeList => false | id_isNePal,id_a__U71 => false | id_isNePal,id_U71 => false | id_isNePal,id_a__U41 => false | id_isNePal,id_U41 => false | id_isNePal,id_a__isPal => false | id_tt,id_a____ => true | id_tt,id_a => true | id_tt,id_a__U42 => true | id_tt,id_U42 => true | id_tt,id_a__U21 => true | id_tt,id_U21 => true | id_tt,id_a__U72 => true | id_tt,id_U72 => true | id_tt,id_a__U11 => true | id_tt,id_u => true | id_tt,id_a__U53 => true | id_tt,id_U53 => true | id_tt,id_a__U31 => true | id_tt,id_U31 => true | id_tt,id_isPalListKind => true | id_tt,id_mark => true | id_tt,id_i => true | id_tt,id_a__U51 => true | id_tt,id_U51 => true | id_tt,id_a__isList => true | id_tt,id_isList => true | id_tt,id_a__and => true | id_tt,id_a__U12 => true | id_tt,id_U12 => true | id_tt,id_a__U62 => true | id_tt,id_U62 => true | id_tt,id_a__isQid => true | id_tt,id_isQid => true | id_tt,id_isPal => true | id_tt,id___ => true | id_tt,id_e => true | id_tt,id_a__U43 => true | id_tt,id_U43 => true | id_tt,id_a__U22 => true | id_tt,id_U22 => true | id_tt,id_a__isNePal => true | id_tt,id_isNePal => true | id_tt,id_tt => true | id_tt,id_U11 => false | id_tt,id_a__U61 => false | id_tt,id_U61 => false | id_tt,id_a__U32 => false | id_tt,id_U32 => false | id_tt,id_and => false | id_tt,id_nil => false | id_tt,id_o => false | id_tt,id_a__U52 => false | id_tt,id_U52 => false | id_tt,id_a__U23 => false | id_tt,id_U23 => false | id_tt,id_a__isPalListKind => false | id_tt,id_a__isNeList => false | id_tt,id_isNeList => false | id_tt,id_a__U71 => false | id_tt,id_U71 => false | id_tt,id_a__U41 => false | id_tt,id_U41 => false | id_tt,id_a__isPal => false | id_U11,id_a____ => true | id_U11,id_a => true | id_U11,id_a__U42 => true | id_U11,id_U42 => true | id_U11,id_a__U21 => true | id_U11,id_U21 => true | id_U11,id_a__U72 => true | id_U11,id_U72 => true | id_U11,id_a__U11 => true | id_U11,id_u => true | id_U11,id_a__U53 => true | id_U11,id_U53 => true | id_U11,id_a__U31 => true | id_U11,id_U31 => true | id_U11,id_isPalListKind => true | id_U11,id_mark => true | id_U11,id_i => true | id_U11,id_a__U51 => true | id_U11,id_U51 => true | id_U11,id_a__isList => true | id_U11,id_isList => true | id_U11,id_a__and => true | id_U11,id_a__U12 => true | id_U11,id_U12 => true | id_U11,id_a__U62 => true | id_U11,id_U62 => true | id_U11,id_a__isQid => true | id_U11,id_isQid => true | id_U11,id_isPal => true | id_U11,id___ => true | id_U11,id_e => true | id_U11,id_a__U43 => true | id_U11,id_U43 => true | id_U11,id_a__U22 => true | id_U11,id_U22 => true | id_U11,id_a__isNePal => true | id_U11,id_isNePal => true | id_U11,id_tt => true | id_U11,id_U11 => true | id_U11,id_a__U61 => false | id_U11,id_U61 => false | id_U11,id_a__U32 => false | id_U11,id_U32 => false | id_U11,id_and => false | id_U11,id_nil => false | id_U11,id_o => false | id_U11,id_a__U52 => false | id_U11,id_U52 => false | id_U11,id_a__U23 => false | id_U11,id_U23 => false | id_U11,id_a__isPalListKind => false | id_U11,id_a__isNeList => false | id_U11,id_isNeList => false | id_U11,id_a__U71 => false | id_U11,id_U71 => false | id_U11,id_a__U41 => false | id_U11,id_U41 => false | id_U11,id_a__isPal => false | id_a__U61,id_a____ => true | id_a__U61,id_a => true | id_a__U61,id_a__U42 => true | id_a__U61,id_U42 => true | id_a__U61,id_a__U21 => true | id_a__U61,id_U21 => true | id_a__U61,id_a__U72 => true | id_a__U61,id_U72 => true | id_a__U61,id_a__U11 => true | id_a__U61,id_u => true | id_a__U61,id_a__U53 => true | id_a__U61,id_U53 => true | id_a__U61,id_a__U31 => true | id_a__U61,id_U31 => true | id_a__U61,id_isPalListKind => true | id_a__U61,id_mark => true | id_a__U61,id_i => true | id_a__U61,id_a__U51 => true | id_a__U61,id_U51 => true | id_a__U61,id_a__isList => true | id_a__U61,id_isList => true | id_a__U61,id_a__and => true | id_a__U61,id_a__U12 => true | id_a__U61,id_U12 => true | id_a__U61,id_a__U62 => true | id_a__U61,id_U62 => true | id_a__U61,id_a__isQid => true | id_a__U61,id_isQid => true | id_a__U61,id_isPal => true | id_a__U61,id___ => true | id_a__U61,id_e => true | id_a__U61,id_a__U43 => true | id_a__U61,id_U43 => true | id_a__U61,id_a__U22 => true | id_a__U61,id_U22 => true | id_a__U61,id_a__isNePal => true | id_a__U61,id_isNePal => true | id_a__U61,id_tt => true | id_a__U61,id_U11 => true | id_a__U61,id_a__U61 => true | id_a__U61,id_U61 => false | id_a__U61,id_a__U32 => false | id_a__U61,id_U32 => false | id_a__U61,id_and => false | id_a__U61,id_nil => false | id_a__U61,id_o => false | id_a__U61,id_a__U52 => false | id_a__U61,id_U52 => false | id_a__U61,id_a__U23 => false | id_a__U61,id_U23 => false | id_a__U61,id_a__isPalListKind => false | id_a__U61,id_a__isNeList => false | id_a__U61,id_isNeList => false | id_a__U61,id_a__U71 => false | id_a__U61,id_U71 => false | id_a__U61,id_a__U41 => false | id_a__U61,id_U41 => false | id_a__U61,id_a__isPal => false | id_U61,id_a____ => true | id_U61,id_a => true | id_U61,id_a__U42 => true | id_U61,id_U42 => true | id_U61,id_a__U21 => true | id_U61,id_U21 => true | id_U61,id_a__U72 => true | id_U61,id_U72 => true | id_U61,id_a__U11 => true | id_U61,id_u => true | id_U61,id_a__U53 => true | id_U61,id_U53 => true | id_U61,id_a__U31 => true | id_U61,id_U31 => true | id_U61,id_isPalListKind => true | id_U61,id_mark => true | id_U61,id_i => true | id_U61,id_a__U51 => true | id_U61,id_U51 => true | id_U61,id_a__isList => true | id_U61,id_isList => true | id_U61,id_a__and => true | id_U61,id_a__U12 => true | id_U61,id_U12 => true | id_U61,id_a__U62 => true | id_U61,id_U62 => true | id_U61,id_a__isQid => true | id_U61,id_isQid => true | id_U61,id_isPal => true | id_U61,id___ => true | id_U61,id_e => true | id_U61,id_a__U43 => true | id_U61,id_U43 => true | id_U61,id_a__U22 => true | id_U61,id_U22 => true | id_U61,id_a__isNePal => true | id_U61,id_isNePal => true | id_U61,id_tt => true | id_U61,id_U11 => true | id_U61,id_a__U61 => true | id_U61,id_U61 => true | id_U61,id_a__U32 => false | id_U61,id_U32 => false | id_U61,id_and => false | id_U61,id_nil => false | id_U61,id_o => false | id_U61,id_a__U52 => false | id_U61,id_U52 => false | id_U61,id_a__U23 => false | id_U61,id_U23 => false | id_U61,id_a__isPalListKind => false | id_U61,id_a__isNeList => false | id_U61,id_isNeList => false | id_U61,id_a__U71 => false | id_U61,id_U71 => false | id_U61,id_a__U41 => false | id_U61,id_U41 => false | id_U61,id_a__isPal => false | id_a__U32,id_a____ => true | id_a__U32,id_a => true | id_a__U32,id_a__U42 => true | id_a__U32,id_U42 => true | id_a__U32,id_a__U21 => true | id_a__U32,id_U21 => true | id_a__U32,id_a__U72 => true | id_a__U32,id_U72 => true | id_a__U32,id_a__U11 => true | id_a__U32,id_u => true | id_a__U32,id_a__U53 => true | id_a__U32,id_U53 => true | id_a__U32,id_a__U31 => true | id_a__U32,id_U31 => true | id_a__U32,id_isPalListKind => true | id_a__U32,id_mark => true | id_a__U32,id_i => true | id_a__U32,id_a__U51 => true | id_a__U32,id_U51 => true | id_a__U32,id_a__isList => true | id_a__U32,id_isList => true | id_a__U32,id_a__and => true | id_a__U32,id_a__U12 => true | id_a__U32,id_U12 => true | id_a__U32,id_a__U62 => true | id_a__U32,id_U62 => true | id_a__U32,id_a__isQid => true | id_a__U32,id_isQid => true | id_a__U32,id_isPal => true | id_a__U32,id___ => true | id_a__U32,id_e => true | id_a__U32,id_a__U43 => true | id_a__U32,id_U43 => true | id_a__U32,id_a__U22 => true | id_a__U32,id_U22 => true | id_a__U32,id_a__isNePal => true | id_a__U32,id_isNePal => true | id_a__U32,id_tt => true | id_a__U32,id_U11 => true | id_a__U32,id_a__U61 => true | id_a__U32,id_U61 => true | id_a__U32,id_a__U32 => true | id_a__U32,id_U32 => false | id_a__U32,id_and => false | id_a__U32,id_nil => false | id_a__U32,id_o => false | id_a__U32,id_a__U52 => false | id_a__U32,id_U52 => false | id_a__U32,id_a__U23 => false | id_a__U32,id_U23 => false | id_a__U32,id_a__isPalListKind => false | id_a__U32,id_a__isNeList => false | id_a__U32,id_isNeList => false | id_a__U32,id_a__U71 => false | id_a__U32,id_U71 => false | id_a__U32,id_a__U41 => false | id_a__U32,id_U41 => false | id_a__U32,id_a__isPal => false | id_U32,id_a____ => true | id_U32,id_a => true | id_U32,id_a__U42 => true | id_U32,id_U42 => true | id_U32,id_a__U21 => true | id_U32,id_U21 => true | id_U32,id_a__U72 => true | id_U32,id_U72 => true | id_U32,id_a__U11 => true | id_U32,id_u => true | id_U32,id_a__U53 => true | id_U32,id_U53 => true | id_U32,id_a__U31 => true | id_U32,id_U31 => true | id_U32,id_isPalListKind => true | id_U32,id_mark => true | id_U32,id_i => true | id_U32,id_a__U51 => true | id_U32,id_U51 => true | id_U32,id_a__isList => true | id_U32,id_isList => true | id_U32,id_a__and => true | id_U32,id_a__U12 => true | id_U32,id_U12 => true | id_U32,id_a__U62 => true | id_U32,id_U62 => true | id_U32,id_a__isQid => true | id_U32,id_isQid => true | id_U32,id_isPal => true | id_U32,id___ => true | id_U32,id_e => true | id_U32,id_a__U43 => true | id_U32,id_U43 => true | id_U32,id_a__U22 => true | id_U32,id_U22 => true | id_U32,id_a__isNePal => true | id_U32,id_isNePal => true | id_U32,id_tt => true | id_U32,id_U11 => true | id_U32,id_a__U61 => true | id_U32,id_U61 => true | id_U32,id_a__U32 => true | id_U32,id_U32 => true | id_U32,id_and => false | id_U32,id_nil => false | id_U32,id_o => false | id_U32,id_a__U52 => false | id_U32,id_U52 => false | id_U32,id_a__U23 => false | id_U32,id_U23 => false | id_U32,id_a__isPalListKind => false | id_U32,id_a__isNeList => false | id_U32,id_isNeList => false | id_U32,id_a__U71 => false | id_U32,id_U71 => false | id_U32,id_a__U41 => false | id_U32,id_U41 => false | id_U32,id_a__isPal => false | id_and,id_a____ => true | id_and,id_a => true | id_and,id_a__U42 => true | id_and,id_U42 => true | id_and,id_a__U21 => true | id_and,id_U21 => true | id_and,id_a__U72 => true | id_and,id_U72 => true | id_and,id_a__U11 => true | id_and,id_u => true | id_and,id_a__U53 => true | id_and,id_U53 => true | id_and,id_a__U31 => true | id_and,id_U31 => true | id_and,id_isPalListKind => true | id_and,id_mark => true | id_and,id_i => true | id_and,id_a__U51 => true | id_and,id_U51 => true | id_and,id_a__isList => true | id_and,id_isList => true | id_and,id_a__and => true | id_and,id_a__U12 => true | id_and,id_U12 => true | id_and,id_a__U62 => true | id_and,id_U62 => true | id_and,id_a__isQid => true | id_and,id_isQid => true | id_and,id_isPal => true | id_and,id___ => true | id_and,id_e => true | id_and,id_a__U43 => true | id_and,id_U43 => true | id_and,id_a__U22 => true | id_and,id_U22 => true | id_and,id_a__isNePal => true | id_and,id_isNePal => true | id_and,id_tt => true | id_and,id_U11 => true | id_and,id_a__U61 => true | id_and,id_U61 => true | id_and,id_a__U32 => true | id_and,id_U32 => true | id_and,id_and => true | id_and,id_nil => false | id_and,id_o => false | id_and,id_a__U52 => false | id_and,id_U52 => false | id_and,id_a__U23 => false | id_and,id_U23 => false | id_and,id_a__isPalListKind => false | id_and,id_a__isNeList => false | id_and,id_isNeList => false | id_and,id_a__U71 => false | id_and,id_U71 => false | id_and,id_a__U41 => false | id_and,id_U41 => false | id_and,id_a__isPal => false | id_nil,id_a____ => true | id_nil,id_a => true | id_nil,id_a__U42 => true | id_nil,id_U42 => true | id_nil,id_a__U21 => true | id_nil,id_U21 => true | id_nil,id_a__U72 => true | id_nil,id_U72 => true | id_nil,id_a__U11 => true | id_nil,id_u => true | id_nil,id_a__U53 => true | id_nil,id_U53 => true | id_nil,id_a__U31 => true | id_nil,id_U31 => true | id_nil,id_isPalListKind => true | id_nil,id_mark => true | id_nil,id_i => true | id_nil,id_a__U51 => true | id_nil,id_U51 => true | id_nil,id_a__isList => true | id_nil,id_isList => true | id_nil,id_a__and => true | id_nil,id_a__U12 => true | id_nil,id_U12 => true | id_nil,id_a__U62 => true | id_nil,id_U62 => true | id_nil,id_a__isQid => true | id_nil,id_isQid => true | id_nil,id_isPal => true | id_nil,id___ => true | id_nil,id_e => true | id_nil,id_a__U43 => true | id_nil,id_U43 => true | id_nil,id_a__U22 => true | id_nil,id_U22 => true | id_nil,id_a__isNePal => true | id_nil,id_isNePal => true | id_nil,id_tt => true | id_nil,id_U11 => true | id_nil,id_a__U61 => true | id_nil,id_U61 => true | id_nil,id_a__U32 => true | id_nil,id_U32 => true | id_nil,id_and => true | id_nil,id_nil => true | id_nil,id_o => false | id_nil,id_a__U52 => false | id_nil,id_U52 => false | id_nil,id_a__U23 => false | id_nil,id_U23 => false | id_nil,id_a__isPalListKind => false | id_nil,id_a__isNeList => false | id_nil,id_isNeList => false | id_nil,id_a__U71 => false | id_nil,id_U71 => false | id_nil,id_a__U41 => false | id_nil,id_U41 => false | id_nil,id_a__isPal => false | id_o,id_a____ => true | id_o,id_a => true | id_o,id_a__U42 => true | id_o,id_U42 => true | id_o,id_a__U21 => true | id_o,id_U21 => true | id_o,id_a__U72 => true | id_o,id_U72 => true | id_o,id_a__U11 => true | id_o,id_u => true | id_o,id_a__U53 => true | id_o,id_U53 => true | id_o,id_a__U31 => true | id_o,id_U31 => true | id_o,id_isPalListKind => true | id_o,id_mark => true | id_o,id_i => true | id_o,id_a__U51 => true | id_o,id_U51 => true | id_o,id_a__isList => true | id_o,id_isList => true | id_o,id_a__and => true | id_o,id_a__U12 => true | id_o,id_U12 => true | id_o,id_a__U62 => true | id_o,id_U62 => true | id_o,id_a__isQid => true | id_o,id_isQid => true | id_o,id_isPal => true | id_o,id___ => true | id_o,id_e => true | id_o,id_a__U43 => true | id_o,id_U43 => true | id_o,id_a__U22 => true | id_o,id_U22 => true | id_o,id_a__isNePal => true | id_o,id_isNePal => true | id_o,id_tt => true | id_o,id_U11 => true | id_o,id_a__U61 => true | id_o,id_U61 => true | id_o,id_a__U32 => true | id_o,id_U32 => true | id_o,id_and => true | id_o,id_nil => true | id_o,id_o => true | id_o,id_a__U52 => false | id_o,id_U52 => false | id_o,id_a__U23 => false | id_o,id_U23 => false | id_o,id_a__isPalListKind => false | id_o,id_a__isNeList => false | id_o,id_isNeList => false | id_o,id_a__U71 => false | id_o,id_U71 => false | id_o,id_a__U41 => false | id_o,id_U41 => false | id_o,id_a__isPal => false | id_a__U52,id_a____ => true | id_a__U52,id_a => true | id_a__U52,id_a__U42 => true | id_a__U52,id_U42 => true | id_a__U52,id_a__U21 => true | id_a__U52,id_U21 => true | id_a__U52,id_a__U72 => true | id_a__U52,id_U72 => true | id_a__U52,id_a__U11 => true | id_a__U52,id_u => true | id_a__U52,id_a__U53 => true | id_a__U52,id_U53 => true | id_a__U52,id_a__U31 => true | id_a__U52,id_U31 => true | id_a__U52,id_isPalListKind => true | id_a__U52,id_mark => true | id_a__U52,id_i => true | id_a__U52,id_a__U51 => true | id_a__U52,id_U51 => true | id_a__U52,id_a__isList => true | id_a__U52,id_isList => true | id_a__U52,id_a__and => true | id_a__U52,id_a__U12 => true | id_a__U52,id_U12 => true | id_a__U52,id_a__U62 => true | id_a__U52,id_U62 => true | id_a__U52,id_a__isQid => true | id_a__U52,id_isQid => true | id_a__U52,id_isPal => true | id_a__U52,id___ => true | id_a__U52,id_e => true | id_a__U52,id_a__U43 => true | id_a__U52,id_U43 => true | id_a__U52,id_a__U22 => true | id_a__U52,id_U22 => true | id_a__U52,id_a__isNePal => true | id_a__U52,id_isNePal => true | id_a__U52,id_tt => true | id_a__U52,id_U11 => true | id_a__U52,id_a__U61 => true | id_a__U52,id_U61 => true | id_a__U52,id_a__U32 => true | id_a__U52,id_U32 => true | id_a__U52,id_and => true | id_a__U52,id_nil => true | id_a__U52,id_o => true | id_a__U52,id_a__U52 => true | id_a__U52,id_U52 => false | id_a__U52,id_a__U23 => false | id_a__U52,id_U23 => false | id_a__U52,id_a__isPalListKind => false | id_a__U52,id_a__isNeList => false | id_a__U52,id_isNeList => false | id_a__U52,id_a__U71 => false | id_a__U52,id_U71 => false | id_a__U52,id_a__U41 => false | id_a__U52,id_U41 => false | id_a__U52,id_a__isPal => false | id_U52,id_a____ => true | id_U52,id_a => true | id_U52,id_a__U42 => true | id_U52,id_U42 => true | id_U52,id_a__U21 => true | id_U52,id_U21 => true | id_U52,id_a__U72 => true | id_U52,id_U72 => true | id_U52,id_a__U11 => true | id_U52,id_u => true | id_U52,id_a__U53 => true | id_U52,id_U53 => true | id_U52,id_a__U31 => true | id_U52,id_U31 => true | id_U52,id_isPalListKind => true | id_U52,id_mark => true | id_U52,id_i => true | id_U52,id_a__U51 => true | id_U52,id_U51 => true | id_U52,id_a__isList => true | id_U52,id_isList => true | id_U52,id_a__and => true | id_U52,id_a__U12 => true | id_U52,id_U12 => true | id_U52,id_a__U62 => true | id_U52,id_U62 => true | id_U52,id_a__isQid => true | id_U52,id_isQid => true | id_U52,id_isPal => true | id_U52,id___ => true | id_U52,id_e => true | id_U52,id_a__U43 => true | id_U52,id_U43 => true | id_U52,id_a__U22 => true | id_U52,id_U22 => true | id_U52,id_a__isNePal => true | id_U52,id_isNePal => true | id_U52,id_tt => true | id_U52,id_U11 => true | id_U52,id_a__U61 => true | id_U52,id_U61 => true | id_U52,id_a__U32 => true | id_U52,id_U32 => true | id_U52,id_and => true | id_U52,id_nil => true | id_U52,id_o => true | id_U52,id_a__U52 => true | id_U52,id_U52 => true | id_U52,id_a__U23 => false | id_U52,id_U23 => false | id_U52,id_a__isPalListKind => false | id_U52,id_a__isNeList => false | id_U52,id_isNeList => false | id_U52,id_a__U71 => false | id_U52,id_U71 => false | id_U52,id_a__U41 => false | id_U52,id_U41 => false | id_U52,id_a__isPal => false | id_a__U23,id_a____ => true | id_a__U23,id_a => true | id_a__U23,id_a__U42 => true | id_a__U23,id_U42 => true | id_a__U23,id_a__U21 => true | id_a__U23,id_U21 => true | id_a__U23,id_a__U72 => true | id_a__U23,id_U72 => true | id_a__U23,id_a__U11 => true | id_a__U23,id_u => true | id_a__U23,id_a__U53 => true | id_a__U23,id_U53 => true | id_a__U23,id_a__U31 => true | id_a__U23,id_U31 => true | id_a__U23,id_isPalListKind => true | id_a__U23,id_mark => true | id_a__U23,id_i => true | id_a__U23,id_a__U51 => true | id_a__U23,id_U51 => true | id_a__U23,id_a__isList => true | id_a__U23,id_isList => true | id_a__U23,id_a__and => true | id_a__U23,id_a__U12 => true | id_a__U23,id_U12 => true | id_a__U23,id_a__U62 => true | id_a__U23,id_U62 => true | id_a__U23,id_a__isQid => true | id_a__U23,id_isQid => true | id_a__U23,id_isPal => true | id_a__U23,id___ => true | id_a__U23,id_e => true | id_a__U23,id_a__U43 => true | id_a__U23,id_U43 => true | id_a__U23,id_a__U22 => true | id_a__U23,id_U22 => true | id_a__U23,id_a__isNePal => true | id_a__U23,id_isNePal => true | id_a__U23,id_tt => true | id_a__U23,id_U11 => true | id_a__U23,id_a__U61 => true | id_a__U23,id_U61 => true | id_a__U23,id_a__U32 => true | id_a__U23,id_U32 => true | id_a__U23,id_and => true | id_a__U23,id_nil => true | id_a__U23,id_o => true | id_a__U23,id_a__U52 => true | id_a__U23,id_U52 => true | id_a__U23,id_a__U23 => true | id_a__U23,id_U23 => false | id_a__U23,id_a__isPalListKind => false | id_a__U23,id_a__isNeList => false | id_a__U23,id_isNeList => false | id_a__U23,id_a__U71 => false | id_a__U23,id_U71 => false | id_a__U23,id_a__U41 => false | id_a__U23,id_U41 => false | id_a__U23,id_a__isPal => false | id_U23,id_a____ => true | id_U23,id_a => true | id_U23,id_a__U42 => true | id_U23,id_U42 => true | id_U23,id_a__U21 => true | id_U23,id_U21 => true | id_U23,id_a__U72 => true | id_U23,id_U72 => true | id_U23,id_a__U11 => true | id_U23,id_u => true | id_U23,id_a__U53 => true | id_U23,id_U53 => true | id_U23,id_a__U31 => true | id_U23,id_U31 => true | id_U23,id_isPalListKind => true | id_U23,id_mark => true | id_U23,id_i => true | id_U23,id_a__U51 => true | id_U23,id_U51 => true | id_U23,id_a__isList => true | id_U23,id_isList => true | id_U23,id_a__and => true | id_U23,id_a__U12 => true | id_U23,id_U12 => true | id_U23,id_a__U62 => true | id_U23,id_U62 => true | id_U23,id_a__isQid => true | id_U23,id_isQid => true | id_U23,id_isPal => true | id_U23,id___ => true | id_U23,id_e => true | id_U23,id_a__U43 => true | id_U23,id_U43 => true | id_U23,id_a__U22 => true | id_U23,id_U22 => true | id_U23,id_a__isNePal => true | id_U23,id_isNePal => true | id_U23,id_tt => true | id_U23,id_U11 => true | id_U23,id_a__U61 => true | id_U23,id_U61 => true | id_U23,id_a__U32 => true | id_U23,id_U32 => true | id_U23,id_and => true | id_U23,id_nil => true | id_U23,id_o => true | id_U23,id_a__U52 => true | id_U23,id_U52 => true | id_U23,id_a__U23 => true | id_U23,id_U23 => true | id_U23,id_a__isPalListKind => false | id_U23,id_a__isNeList => false | id_U23,id_isNeList => false | id_U23,id_a__U71 => false | id_U23,id_U71 => false | id_U23,id_a__U41 => false | id_U23,id_U41 => false | id_U23,id_a__isPal => false | id_a__isPalListKind,id_a____ => true | id_a__isPalListKind,id_a => true | id_a__isPalListKind,id_a__U42 => true | id_a__isPalListKind,id_U42 => true | id_a__isPalListKind,id_a__U21 => true | id_a__isPalListKind,id_U21 => true | id_a__isPalListKind,id_a__U72 => true | id_a__isPalListKind,id_U72 => true | id_a__isPalListKind,id_a__U11 => true | id_a__isPalListKind,id_u => true | id_a__isPalListKind,id_a__U53 => true | id_a__isPalListKind,id_U53 => true | id_a__isPalListKind,id_a__U31 => true | id_a__isPalListKind,id_U31 => true | id_a__isPalListKind,id_isPalListKind => true | id_a__isPalListKind,id_mark => true | id_a__isPalListKind,id_i => true | id_a__isPalListKind,id_a__U51 => true | id_a__isPalListKind,id_U51 => true | id_a__isPalListKind,id_a__isList => true | id_a__isPalListKind,id_isList => true | id_a__isPalListKind,id_a__and => true | id_a__isPalListKind,id_a__U12 => true | id_a__isPalListKind,id_U12 => true | id_a__isPalListKind,id_a__U62 => true | id_a__isPalListKind,id_U62 => true | id_a__isPalListKind,id_a__isQid => true | id_a__isPalListKind,id_isQid => true | id_a__isPalListKind,id_isPal => true | id_a__isPalListKind,id___ => true | id_a__isPalListKind,id_e => true | id_a__isPalListKind,id_a__U43 => true | id_a__isPalListKind,id_U43 => true | id_a__isPalListKind,id_a__U22 => true | id_a__isPalListKind,id_U22 => true | id_a__isPalListKind,id_a__isNePal => true | id_a__isPalListKind,id_isNePal => true | id_a__isPalListKind,id_tt => true | id_a__isPalListKind,id_U11 => true | id_a__isPalListKind,id_a__U61 => true | id_a__isPalListKind,id_U61 => true | id_a__isPalListKind,id_a__U32 => true | id_a__isPalListKind,id_U32 => true | id_a__isPalListKind,id_and => true | id_a__isPalListKind,id_nil => true | id_a__isPalListKind,id_o => true | id_a__isPalListKind,id_a__U52 => true | id_a__isPalListKind,id_U52 => true | id_a__isPalListKind,id_a__U23 => true | id_a__isPalListKind,id_U23 => true | id_a__isPalListKind,id_a__isPalListKind => true | id_a__isPalListKind,id_a__isNeList => false | id_a__isPalListKind,id_isNeList => false | id_a__isPalListKind,id_a__U71 => false | id_a__isPalListKind,id_U71 => false | id_a__isPalListKind,id_a__U41 => false | id_a__isPalListKind,id_U41 => false | id_a__isPalListKind,id_a__isPal => false | id_a__isNeList,id_a____ => true | id_a__isNeList,id_a => true | id_a__isNeList,id_a__U42 => true | id_a__isNeList,id_U42 => true | id_a__isNeList,id_a__U21 => true | id_a__isNeList,id_U21 => true | id_a__isNeList,id_a__U72 => true | id_a__isNeList,id_U72 => true | id_a__isNeList,id_a__U11 => true | id_a__isNeList,id_u => true | id_a__isNeList,id_a__U53 => true | id_a__isNeList,id_U53 => true | id_a__isNeList,id_a__U31 => true | id_a__isNeList,id_U31 => true | id_a__isNeList,id_isPalListKind => true | id_a__isNeList,id_mark => true | id_a__isNeList,id_i => true | id_a__isNeList,id_a__U51 => true | id_a__isNeList,id_U51 => true | id_a__isNeList,id_a__isList => true | id_a__isNeList,id_isList => true | id_a__isNeList,id_a__and => true | id_a__isNeList,id_a__U12 => true | id_a__isNeList,id_U12 => true | id_a__isNeList,id_a__U62 => true | id_a__isNeList,id_U62 => true | id_a__isNeList,id_a__isQid => true | id_a__isNeList,id_isQid => true | id_a__isNeList,id_isPal => true | id_a__isNeList,id___ => true | id_a__isNeList,id_e => true | id_a__isNeList,id_a__U43 => true | id_a__isNeList,id_U43 => true | id_a__isNeList,id_a__U22 => true | id_a__isNeList,id_U22 => true | id_a__isNeList,id_a__isNePal => true | id_a__isNeList,id_isNePal => true | id_a__isNeList,id_tt => true | id_a__isNeList,id_U11 => true | id_a__isNeList,id_a__U61 => true | id_a__isNeList,id_U61 => true | id_a__isNeList,id_a__U32 => true | id_a__isNeList,id_U32 => true | id_a__isNeList,id_and => true | id_a__isNeList,id_nil => true | id_a__isNeList,id_o => true | id_a__isNeList,id_a__U52 => true | id_a__isNeList,id_U52 => true | id_a__isNeList,id_a__U23 => true | id_a__isNeList,id_U23 => true | id_a__isNeList,id_a__isPalListKind => true | id_a__isNeList,id_a__isNeList => true | id_a__isNeList,id_isNeList => false | id_a__isNeList,id_a__U71 => false | id_a__isNeList,id_U71 => false | id_a__isNeList,id_a__U41 => false | id_a__isNeList,id_U41 => false | id_a__isNeList,id_a__isPal => false | id_isNeList,id_a____ => true | id_isNeList,id_a => true | id_isNeList,id_a__U42 => true | id_isNeList,id_U42 => true | id_isNeList,id_a__U21 => true | id_isNeList,id_U21 => true | id_isNeList,id_a__U72 => true | id_isNeList,id_U72 => true | id_isNeList,id_a__U11 => true | id_isNeList,id_u => true | id_isNeList,id_a__U53 => true | id_isNeList,id_U53 => true | id_isNeList,id_a__U31 => true | id_isNeList,id_U31 => true | id_isNeList,id_isPalListKind => true | id_isNeList,id_mark => true | id_isNeList,id_i => true | id_isNeList,id_a__U51 => true | id_isNeList,id_U51 => true | id_isNeList,id_a__isList => true | id_isNeList,id_isList => true | id_isNeList,id_a__and => true | id_isNeList,id_a__U12 => true | id_isNeList,id_U12 => true | id_isNeList,id_a__U62 => true | id_isNeList,id_U62 => true | id_isNeList,id_a__isQid => true | id_isNeList,id_isQid => true | id_isNeList,id_isPal => true | id_isNeList,id___ => true | id_isNeList,id_e => true | id_isNeList,id_a__U43 => true | id_isNeList,id_U43 => true | id_isNeList,id_a__U22 => true | id_isNeList,id_U22 => true | id_isNeList,id_a__isNePal => true | id_isNeList,id_isNePal => true | id_isNeList,id_tt => true | id_isNeList,id_U11 => true | id_isNeList,id_a__U61 => true | id_isNeList,id_U61 => true | id_isNeList,id_a__U32 => true | id_isNeList,id_U32 => true | id_isNeList,id_and => true | id_isNeList,id_nil => true | id_isNeList,id_o => true | id_isNeList,id_a__U52 => true | id_isNeList,id_U52 => true | id_isNeList,id_a__U23 => true | id_isNeList,id_U23 => true | id_isNeList,id_a__isPalListKind => true | id_isNeList,id_a__isNeList => true | id_isNeList,id_isNeList => true | id_isNeList,id_a__U71 => false | id_isNeList,id_U71 => false | id_isNeList,id_a__U41 => false | id_isNeList,id_U41 => false | id_isNeList,id_a__isPal => false | id_a__U71,id_a____ => true | id_a__U71,id_a => true | id_a__U71,id_a__U42 => true | id_a__U71,id_U42 => true | id_a__U71,id_a__U21 => true | id_a__U71,id_U21 => true | id_a__U71,id_a__U72 => true | id_a__U71,id_U72 => true | id_a__U71,id_a__U11 => true | id_a__U71,id_u => true | id_a__U71,id_a__U53 => true | id_a__U71,id_U53 => true | id_a__U71,id_a__U31 => true | id_a__U71,id_U31 => true | id_a__U71,id_isPalListKind => true | id_a__U71,id_mark => true | id_a__U71,id_i => true | id_a__U71,id_a__U51 => true | id_a__U71,id_U51 => true | id_a__U71,id_a__isList => true | id_a__U71,id_isList => true | id_a__U71,id_a__and => true | id_a__U71,id_a__U12 => true | id_a__U71,id_U12 => true | id_a__U71,id_a__U62 => true | id_a__U71,id_U62 => true | id_a__U71,id_a__isQid => true | id_a__U71,id_isQid => true | id_a__U71,id_isPal => true | id_a__U71,id___ => true | id_a__U71,id_e => true | id_a__U71,id_a__U43 => true | id_a__U71,id_U43 => true | id_a__U71,id_a__U22 => true | id_a__U71,id_U22 => true | id_a__U71,id_a__isNePal => true | id_a__U71,id_isNePal => true | id_a__U71,id_tt => true | id_a__U71,id_U11 => true | id_a__U71,id_a__U61 => true | id_a__U71,id_U61 => true | id_a__U71,id_a__U32 => true | id_a__U71,id_U32 => true | id_a__U71,id_and => true | id_a__U71,id_nil => true | id_a__U71,id_o => true | id_a__U71,id_a__U52 => true | id_a__U71,id_U52 => true | id_a__U71,id_a__U23 => true | id_a__U71,id_U23 => true | id_a__U71,id_a__isPalListKind => true | id_a__U71,id_a__isNeList => true | id_a__U71,id_isNeList => true | id_a__U71,id_a__U71 => true | id_a__U71,id_U71 => false | id_a__U71,id_a__U41 => false | id_a__U71,id_U41 => false | id_a__U71,id_a__isPal => false | id_U71,id_a____ => true | id_U71,id_a => true | id_U71,id_a__U42 => true | id_U71,id_U42 => true | id_U71,id_a__U21 => true | id_U71,id_U21 => true | id_U71,id_a__U72 => true | id_U71,id_U72 => true | id_U71,id_a__U11 => true | id_U71,id_u => true | id_U71,id_a__U53 => true | id_U71,id_U53 => true | id_U71,id_a__U31 => true | id_U71,id_U31 => true | id_U71,id_isPalListKind => true | id_U71,id_mark => true | id_U71,id_i => true | id_U71,id_a__U51 => true | id_U71,id_U51 => true | id_U71,id_a__isList => true | id_U71,id_isList => true | id_U71,id_a__and => true | id_U71,id_a__U12 => true | id_U71,id_U12 => true | id_U71,id_a__U62 => true | id_U71,id_U62 => true | id_U71,id_a__isQid => true | id_U71,id_isQid => true | id_U71,id_isPal => true | id_U71,id___ => true | id_U71,id_e => true | id_U71,id_a__U43 => true | id_U71,id_U43 => true | id_U71,id_a__U22 => true | id_U71,id_U22 => true | id_U71,id_a__isNePal => true | id_U71,id_isNePal => true | id_U71,id_tt => true | id_U71,id_U11 => true | id_U71,id_a__U61 => true | id_U71,id_U61 => true | id_U71,id_a__U32 => true | id_U71,id_U32 => true | id_U71,id_and => true | id_U71,id_nil => true | id_U71,id_o => true | id_U71,id_a__U52 => true | id_U71,id_U52 => true | id_U71,id_a__U23 => true | id_U71,id_U23 => true | id_U71,id_a__isPalListKind => true | id_U71,id_a__isNeList => true | id_U71,id_isNeList => true | id_U71,id_a__U71 => true | id_U71,id_U71 => true | id_U71,id_a__U41 => false | id_U71,id_U41 => false | id_U71,id_a__isPal => false | id_a__U41,id_a____ => true | id_a__U41,id_a => true | id_a__U41,id_a__U42 => true | id_a__U41,id_U42 => true | id_a__U41,id_a__U21 => true | id_a__U41,id_U21 => true | id_a__U41,id_a__U72 => true | id_a__U41,id_U72 => true | id_a__U41,id_a__U11 => true | id_a__U41,id_u => true | id_a__U41,id_a__U53 => true | id_a__U41,id_U53 => true | id_a__U41,id_a__U31 => true | id_a__U41,id_U31 => true | id_a__U41,id_isPalListKind => true | id_a__U41,id_mark => true | id_a__U41,id_i => true | id_a__U41,id_a__U51 => true | id_a__U41,id_U51 => true | id_a__U41,id_a__isList => true | id_a__U41,id_isList => true | id_a__U41,id_a__and => true | id_a__U41,id_a__U12 => true | id_a__U41,id_U12 => true | id_a__U41,id_a__U62 => true | id_a__U41,id_U62 => true | id_a__U41,id_a__isQid => true | id_a__U41,id_isQid => true | id_a__U41,id_isPal => true | id_a__U41,id___ => true | id_a__U41,id_e => true | id_a__U41,id_a__U43 => true | id_a__U41,id_U43 => true | id_a__U41,id_a__U22 => true | id_a__U41,id_U22 => true | id_a__U41,id_a__isNePal => true | id_a__U41,id_isNePal => true | id_a__U41,id_tt => true | id_a__U41,id_U11 => true | id_a__U41,id_a__U61 => true | id_a__U41,id_U61 => true | id_a__U41,id_a__U32 => true | id_a__U41,id_U32 => true | id_a__U41,id_and => true | id_a__U41,id_nil => true | id_a__U41,id_o => true | id_a__U41,id_a__U52 => true | id_a__U41,id_U52 => true | id_a__U41,id_a__U23 => true | id_a__U41,id_U23 => true | id_a__U41,id_a__isPalListKind => true | id_a__U41,id_a__isNeList => true | id_a__U41,id_isNeList => true | id_a__U41,id_a__U71 => true | id_a__U41,id_U71 => true | id_a__U41,id_a__U41 => true | id_a__U41,id_U41 => false | id_a__U41,id_a__isPal => false | id_U41,id_a____ => true | id_U41,id_a => true | id_U41,id_a__U42 => true | id_U41,id_U42 => true | id_U41,id_a__U21 => true | id_U41,id_U21 => true | id_U41,id_a__U72 => true | id_U41,id_U72 => true | id_U41,id_a__U11 => true | id_U41,id_u => true | id_U41,id_a__U53 => true | id_U41,id_U53 => true | id_U41,id_a__U31 => true | id_U41,id_U31 => true | id_U41,id_isPalListKind => true | id_U41,id_mark => true | id_U41,id_i => true | id_U41,id_a__U51 => true | id_U41,id_U51 => true | id_U41,id_a__isList => true | id_U41,id_isList => true | id_U41,id_a__and => true | id_U41,id_a__U12 => true | id_U41,id_U12 => true | id_U41,id_a__U62 => true | id_U41,id_U62 => true | id_U41,id_a__isQid => true | id_U41,id_isQid => true | id_U41,id_isPal => true | id_U41,id___ => true | id_U41,id_e => true | id_U41,id_a__U43 => true | id_U41,id_U43 => true | id_U41,id_a__U22 => true | id_U41,id_U22 => true | id_U41,id_a__isNePal => true | id_U41,id_isNePal => true | id_U41,id_tt => true | id_U41,id_U11 => true | id_U41,id_a__U61 => true | id_U41,id_U61 => true | id_U41,id_a__U32 => true | id_U41,id_U32 => true | id_U41,id_and => true | id_U41,id_nil => true | id_U41,id_o => true | id_U41,id_a__U52 => true | id_U41,id_U52 => true | id_U41,id_a__U23 => true | id_U41,id_U23 => true | id_U41,id_a__isPalListKind => true | id_U41,id_a__isNeList => true | id_U41,id_isNeList => true | id_U41,id_a__U71 => true | id_U41,id_U71 => true | id_U41,id_a__U41 => true | id_U41,id_U41 => true | id_U41,id_a__isPal => false | id_a__isPal,id_a____ => true | id_a__isPal,id_a => true | id_a__isPal,id_a__U42 => true | id_a__isPal,id_U42 => true | id_a__isPal,id_a__U21 => true | id_a__isPal,id_U21 => true | id_a__isPal,id_a__U72 => true | id_a__isPal,id_U72 => true | id_a__isPal,id_a__U11 => true | id_a__isPal,id_u => true | id_a__isPal,id_a__U53 => true | id_a__isPal,id_U53 => true | id_a__isPal,id_a__U31 => true | id_a__isPal,id_U31 => true | id_a__isPal,id_isPalListKind => true | id_a__isPal,id_mark => true | id_a__isPal,id_i => true | id_a__isPal,id_a__U51 => true | id_a__isPal,id_U51 => true | id_a__isPal,id_a__isList => true | id_a__isPal,id_isList => true | id_a__isPal,id_a__and => true | id_a__isPal,id_a__U12 => true | id_a__isPal,id_U12 => true | id_a__isPal,id_a__U62 => true | id_a__isPal,id_U62 => true | id_a__isPal,id_a__isQid => true | id_a__isPal,id_isQid => true | id_a__isPal,id_isPal => true | id_a__isPal,id___ => true | id_a__isPal,id_e => true | id_a__isPal,id_a__U43 => true | id_a__isPal,id_U43 => true | id_a__isPal,id_a__U22 => true | id_a__isPal,id_U22 => true | id_a__isPal,id_a__isNePal => true | id_a__isPal,id_isNePal => true | id_a__isPal,id_tt => true | id_a__isPal,id_U11 => true | id_a__isPal,id_a__U61 => true | id_a__isPal,id_U61 => true | id_a__isPal,id_a__U32 => true | id_a__isPal,id_U32 => true | id_a__isPal,id_and => true | id_a__isPal,id_nil => true | id_a__isPal,id_o => true | id_a__isPal,id_a__U52 => true | id_a__isPal,id_U52 => true | id_a__isPal,id_a__U23 => true | id_a__isPal,id_U23 => true | id_a__isPal,id_a__isPalListKind => true | id_a__isPal,id_a__isNeList => true | id_a__isPal,id_isNeList => true | id_a__isPal,id_a__U71 => true | id_a__isPal,id_U71 => true | id_a__isPal,id_a__U41 => true | id_a__isPal,id_U41 => true | id_a__isPal,id_a__isPal => 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____(__(X_,Y_),Z_) -> a____(mark(X_),a____(mark(Y_),mark(Z_))) *) | R_xml_0_rule_0 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil)):: (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 2)::nil)):: (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 3)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil)):: (algebra.Alg.Var 3)::nil)) (* a____(X_,nil) -> mark(X_) *) | R_xml_0_rule_1 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* a____(nil,X_) -> mark(X_) *) | R_xml_0_rule_2 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Var 1)::nil)) (* a__U11(tt,V_) -> a__U12(a__isNeList(V_)) *) | R_xml_0_rule_3 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_a__isNeList ((algebra.Alg.Var 4)::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 4)::nil)) (* a__U12(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__U12 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)) (* a__U21(tt,V1_,V2_) -> a__U22(a__isList(V1_),V2_) *) | R_xml_0_rule_5 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_a__isList ((algebra.Alg.Var 5)::nil)):: (algebra.Alg.Var 6)::nil)) (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil)) (* a__U22(tt,V2_) -> a__U23(a__isList(V2_)) *) | R_xml_0_rule_6 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U23 ((algebra.Alg.Term algebra.F.id_a__isList ((algebra.Alg.Var 6)::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 6)::nil)) (* a__U23(tt) -> tt *) | R_xml_0_rule_7 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__U23 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)) (* a__U31(tt,V_) -> a__U32(a__isQid(V_)) *) | R_xml_0_rule_8 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Var 4)::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 4)::nil)) (* a__U32(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__U32 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)) (* a__U41(tt,V1_,V2_) -> a__U42(a__isList(V1_),V2_) *) | R_xml_0_rule_10 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isList ((algebra.Alg.Var 5)::nil)):: (algebra.Alg.Var 6)::nil)) (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil)) (* a__U42(tt,V2_) -> a__U43(a__isNeList(V2_)) *) | R_xml_0_rule_11 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_a__isNeList ((algebra.Alg.Var 6)::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 6)::nil)) (* a__U43(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__U43 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)) (* a__U51(tt,V1_,V2_) -> a__U52(a__isNeList(V1_),V2_) *) | R_xml_0_rule_13 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNeList ((algebra.Alg.Var 5)::nil)):: (algebra.Alg.Var 6)::nil)) (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil)) (* a__U52(tt,V2_) -> a__U53(a__isList(V2_)) *) | R_xml_0_rule_14 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_a__isList ((algebra.Alg.Var 6)::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 6)::nil)) (* a__U53(tt) -> 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__U53 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)) (* a__U61(tt,V_) -> a__U62(a__isQid(V_)) *) | R_xml_0_rule_16 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U62 ((algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Var 4)::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 4)::nil)) (* a__U62(tt) -> tt *) | R_xml_0_rule_17 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__U62 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)) (* a__U71(tt,V_) -> a__U72(a__isNePal(V_)) *) | R_xml_0_rule_18 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U72 ((algebra.Alg.Term algebra.F.id_a__isNePal ((algebra.Alg.Var 4)::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 4)::nil)) (* a__U72(tt) -> 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__U72 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)) (* a__and(tt,X_) -> mark(X_) *) | R_xml_0_rule_20 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 1)::nil)) (* a__isList(V_) -> a__U11(a__isPalListKind(V_),V_) *) | R_xml_0_rule_21 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Var 4)::nil))::(algebra.Alg.Var 4)::nil)) (algebra.Alg.Term algebra.F.id_a__isList ((algebra.Alg.Var 4)::nil)) (* a__isList(nil) -> 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__isList ((algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* a__isList(__(V1_,V2_)) -> a__U21(a__and(a__isPalListKind(V1_),isPalListKind(V2_)),V1_,V2_) *) | R_xml_0_rule_23 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind ((algebra.Alg.Var 6)::nil))::nil))::(algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil)) (algebra.Alg.Term algebra.F.id_a__isList ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* a__isNeList(V_) -> a__U31(a__isPalListKind(V_),V_) *) | R_xml_0_rule_24 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Var 4)::nil))::(algebra.Alg.Var 4)::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList ((algebra.Alg.Var 4)::nil)) (* a__isNeList(__(V1_,V2_)) -> a__U41(a__and(a__isPalListKind(V1_),isPalListKind(V2_)),V1_,V2_) *) | R_xml_0_rule_25 : 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__isPalListKind ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind ((algebra.Alg.Var 6)::nil))::nil))::(algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* a__isNeList(__(V1_,V2_)) -> a__U51(a__and(a__isPalListKind(V1_),isPalListKind(V2_)),V1_,V2_) *) | R_xml_0_rule_26 : 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__isPalListKind ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind ((algebra.Alg.Var 6)::nil))::nil))::(algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* a__isNePal(V_) -> a__U61(a__isPalListKind(V_),V_) *) | R_xml_0_rule_27 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Var 4)::nil))::(algebra.Alg.Var 4)::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal ((algebra.Alg.Var 4)::nil)) (* a__isNePal(__(I_,__(P_,I_))) -> a__and(a__and(a__isQid(I_),isPalListKind(I_)), and(isPal(P_),isPalListKind(P_))) *) | R_xml_0_rule_28 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Var 7)::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind ((algebra.Alg.Var 7)::nil))::nil))::(algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isPal ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind ((algebra.Alg.Var 8)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 7)::(algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 8):: (algebra.Alg.Var 7)::nil))::nil))::nil)) (* a__isPal(V_) -> a__U71(a__isPalListKind(V_),V_) *) | R_xml_0_rule_29 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Var 4)::nil))::(algebra.Alg.Var 4)::nil)) (algebra.Alg.Term algebra.F.id_a__isPal ((algebra.Alg.Var 4)::nil)) (* a__isPal(nil) -> tt *) | R_xml_0_rule_30 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isPal ((algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* a__isPalListKind(a) -> tt *) | R_xml_0_rule_31 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Term algebra.F.id_a nil)::nil)) (* a__isPalListKind(e) -> tt *) | R_xml_0_rule_32 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Term algebra.F.id_e nil)::nil)) (* a__isPalListKind(i) -> tt *) | R_xml_0_rule_33 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Term algebra.F.id_i nil)::nil)) (* a__isPalListKind(nil) -> tt *) | R_xml_0_rule_34 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* a__isPalListKind(o) -> tt *) | R_xml_0_rule_35 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Term algebra.F.id_o nil)::nil)) (* a__isPalListKind(u) -> tt *) | R_xml_0_rule_36 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Term algebra.F.id_u nil)::nil)) (* a__isPalListKind(__(V1_,V2_)) -> a__and(a__isPalListKind(V1_),isPalListKind(V2_)) *) | R_xml_0_rule_37 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind ((algebra.Alg.Var 6)::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* a__isQid(a) -> tt *) | R_xml_0_rule_38 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Term algebra.F.id_a nil)::nil)) (* a__isQid(e) -> tt *) | R_xml_0_rule_39 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Term algebra.F.id_e nil)::nil)) (* a__isQid(i) -> tt *) | R_xml_0_rule_40 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Term algebra.F.id_i nil)::nil)) (* a__isQid(o) -> tt *) | R_xml_0_rule_41 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Term algebra.F.id_o nil)::nil)) (* a__isQid(u) -> tt *) | R_xml_0_rule_42 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tt nil) (algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Term algebra.F.id_u nil)::nil)) (* mark(__(X1_,X2_)) -> a____(mark(X1_),mark(X2_)) *) | R_xml_0_rule_43 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 10)::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)) (* mark(U11(X1_,X2_)) -> a__U11(mark(X1_),X2_) *) | R_xml_0_rule_44 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)) (* mark(U12(X_)) -> a__U12(mark(X_)) *) | R_xml_0_rule_45 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U12 ((algebra.Alg.Var 1)::nil))::nil)) (* mark(isNeList(X_)) -> a__isNeList(X_) *) | R_xml_0_rule_46 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__isNeList ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Var 1)::nil))::nil)) (* mark(U21(X1_,X2_,X3_)) -> a__U21(mark(X1_),X2_,X3_) *) | R_xml_0_rule_47 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::(algebra.Alg.Var 11)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10):: (algebra.Alg.Var 11)::nil))::nil)) (* mark(U22(X1_,X2_)) -> a__U22(mark(X1_),X2_) *) | R_xml_0_rule_48 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)) (* mark(isList(X_)) -> a__isList(X_) *) | R_xml_0_rule_49 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__isList ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isList ((algebra.Alg.Var 1)::nil))::nil)) (* mark(U23(X_)) -> a__U23(mark(X_)) *) | R_xml_0_rule_50 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U23 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U23 ((algebra.Alg.Var 1)::nil))::nil)) (* mark(U31(X1_,X2_)) -> a__U31(mark(X1_),X2_) *) | R_xml_0_rule_51 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)) (* mark(U32(X_)) -> a__U32(mark(X_)) *) | R_xml_0_rule_52 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U32 ((algebra.Alg.Var 1)::nil))::nil)) (* mark(isQid(X_)) -> a__isQid(X_) *) | R_xml_0_rule_53 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isQid ((algebra.Alg.Var 1)::nil))::nil)) (* mark(U41(X1_,X2_,X3_)) -> a__U41(mark(X1_),X2_,X3_) *) | R_xml_0_rule_54 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::(algebra.Alg.Var 11)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10):: (algebra.Alg.Var 11)::nil))::nil)) (* mark(U42(X1_,X2_)) -> a__U42(mark(X1_),X2_) *) | R_xml_0_rule_55 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)) (* mark(U43(X_)) -> a__U43(mark(X_)) *) | R_xml_0_rule_56 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U43 ((algebra.Alg.Var 1)::nil))::nil)) (* mark(U51(X1_,X2_,X3_)) -> a__U51(mark(X1_),X2_,X3_) *) | R_xml_0_rule_57 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::(algebra.Alg.Var 11)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10):: (algebra.Alg.Var 11)::nil))::nil)) (* mark(U52(X1_,X2_)) -> a__U52(mark(X1_),X2_) *) | R_xml_0_rule_58 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)) (* mark(U53(X_)) -> a__U53(mark(X_)) *) | R_xml_0_rule_59 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U53 ((algebra.Alg.Var 1)::nil))::nil)) (* mark(U61(X1_,X2_)) -> a__U61(mark(X1_),X2_) *) | R_xml_0_rule_60 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)) (* mark(U62(X_)) -> a__U62(mark(X_)) *) | R_xml_0_rule_61 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U62 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U62 ((algebra.Alg.Var 1)::nil))::nil)) (* mark(U71(X1_,X2_)) -> a__U71(mark(X1_),X2_) *) | R_xml_0_rule_62 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)) (* mark(U72(X_)) -> a__U72(mark(X_)) *) | R_xml_0_rule_63 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__U72 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U72 ((algebra.Alg.Var 1)::nil))::nil)) (* mark(isNePal(X_)) -> a__isNePal(X_) *) | R_xml_0_rule_64 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__isNePal ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isNePal ((algebra.Alg.Var 1)::nil))::nil)) (* mark(and(X1_,X2_)) -> a__and(mark(X1_),X2_) *) | R_xml_0_rule_65 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)) (* mark(isPalListKind(X_)) -> a__isPalListKind(X_) *) | R_xml_0_rule_66 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isPalListKind ((algebra.Alg.Var 1)::nil))::nil)) (* mark(isPal(X_)) -> a__isPal(X_) *) | R_xml_0_rule_67 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a__isPal ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isPal ((algebra.Alg.Var 1)::nil))::nil)) (* mark(nil) -> nil *) | R_xml_0_rule_68 : 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)) (* mark(tt) -> tt *) | R_xml_0_rule_69 : 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(a) -> a *) | R_xml_0_rule_70 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_a nil) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_a nil)::nil)) (* mark(e) -> e *) | R_xml_0_rule_71 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_e nil) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_e nil)::nil)) (* mark(i) -> i *) | R_xml_0_rule_72 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_i nil) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_i nil)::nil)) (* mark(o) -> o *) | R_xml_0_rule_73 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_o nil) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_o nil)::nil)) (* mark(u) -> u *) | R_xml_0_rule_74 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_u nil) (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_u nil)::nil)) (* a____(X1_,X2_) -> __(X1_,X2_) *) | R_xml_0_rule_75 : R_xml_0_rules (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (* a__U11(X1_,X2_) -> U11(X1_,X2_) *) | R_xml_0_rule_76 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (* a__U12(X_) -> U12(X_) *) | R_xml_0_rule_77 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U12 ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Var 1)::nil)) (* a__isNeList(X_) -> isNeList(X_) *) | R_xml_0_rule_78 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList ((algebra.Alg.Var 1)::nil)) (* a__U21(X1_,X2_,X3_) -> U21(X1_,X2_,X3_) *) | R_xml_0_rule_79 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::(algebra.Alg.Var 11)::nil)) (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::(algebra.Alg.Var 11)::nil)) (* a__U22(X1_,X2_) -> U22(X1_,X2_) *) | R_xml_0_rule_80 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (* a__isList(X_) -> isList(X_) *) | R_xml_0_rule_81 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_isList ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__isList ((algebra.Alg.Var 1)::nil)) (* a__U23(X_) -> U23(X_) *) | R_xml_0_rule_82 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U23 ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__U23 ((algebra.Alg.Var 1)::nil)) (* a__U31(X1_,X2_) -> U31(X1_,X2_) *) | R_xml_0_rule_83 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (* a__U32(X_) -> U32(X_) *) | R_xml_0_rule_84 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U32 ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Var 1)::nil)) (* a__isQid(X_) -> isQid(X_) *) | R_xml_0_rule_85 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_isQid ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Var 1)::nil)) (* a__U41(X1_,X2_,X3_) -> U41(X1_,X2_,X3_) *) | R_xml_0_rule_86 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::(algebra.Alg.Var 11)::nil)) (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::(algebra.Alg.Var 11)::nil)) (* a__U42(X1_,X2_) -> U42(X1_,X2_) *) | R_xml_0_rule_87 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (* a__U43(X_) -> U43(X_) *) | R_xml_0_rule_88 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U43 ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Var 1)::nil)) (* a__U51(X1_,X2_,X3_) -> U51(X1_,X2_,X3_) *) | R_xml_0_rule_89 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::(algebra.Alg.Var 11)::nil)) (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::(algebra.Alg.Var 11)::nil)) (* a__U52(X1_,X2_) -> U52(X1_,X2_) *) | R_xml_0_rule_90 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (* a__U53(X_) -> U53(X_) *) | R_xml_0_rule_91 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U53 ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Var 1)::nil)) (* a__U61(X1_,X2_) -> U61(X1_,X2_) *) | R_xml_0_rule_92 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (* a__U62(X_) -> U62(X_) *) | R_xml_0_rule_93 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U62 ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__U62 ((algebra.Alg.Var 1)::nil)) (* a__U71(X1_,X2_) -> U71(X1_,X2_) *) | R_xml_0_rule_94 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (* a__U72(X_) -> U72(X_) *) | R_xml_0_rule_95 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_U72 ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__U72 ((algebra.Alg.Var 1)::nil)) (* a__isNePal(X_) -> isNePal(X_) *) | R_xml_0_rule_96 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_isNePal ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal ((algebra.Alg.Var 1)::nil)) (* a__and(X1_,X2_) -> and(X1_,X2_) *) | R_xml_0_rule_97 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)) (* a__isPalListKind(X_) -> isPalListKind(X_) *) | R_xml_0_rule_98 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_isPalListKind ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Var 1)::nil)) (* a__isPal(X_) -> isPal(X_) *) | R_xml_0_rule_99 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_isPal ((algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_a__isPal ((algebra.Alg.Var 1)::nil)) . Definition R_xml_0_rule_as_list_0 := ((algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil)):: (algebra.Alg.Var 3)::nil)), (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 2)::nil))::(algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 3)::nil))::nil))::nil)))::nil. Definition R_xml_0_rule_as_list_1 := ((algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id_nil nil)::nil)), (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_0. Definition R_xml_0_rule_as_list_2 := ((algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_1. Definition R_xml_0_rule_as_list_3 := ((algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 4)::nil)), (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_a__isNeList ((algebra.Alg.Var 4)::nil))::nil))):: R_xml_0_rule_as_list_2. Definition R_xml_0_rule_as_list_4 := ((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_3. Definition R_xml_0_rule_as_list_5 := ((algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil)), (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_a__isList ((algebra.Alg.Var 5)::nil)):: (algebra.Alg.Var 6)::nil)))::R_xml_0_rule_as_list_4. Definition R_xml_0_rule_as_list_6 := ((algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 6)::nil)), (algebra.Alg.Term algebra.F.id_a__U23 ((algebra.Alg.Term algebra.F.id_a__isList ((algebra.Alg.Var 6)::nil))::nil))):: R_xml_0_rule_as_list_5. Definition R_xml_0_rule_as_list_7 := ((algebra.Alg.Term algebra.F.id_a__U23 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_6. Definition R_xml_0_rule_as_list_8 := ((algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 4)::nil)), (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Var 4)::nil))::nil))):: R_xml_0_rule_as_list_7. Definition R_xml_0_rule_as_list_9 := ((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_8. Definition R_xml_0_rule_as_list_10 := ((algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil)), (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isList ((algebra.Alg.Var 5)::nil)):: (algebra.Alg.Var 6)::nil)))::R_xml_0_rule_as_list_9. Definition R_xml_0_rule_as_list_11 := ((algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 6)::nil)), (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_a__isNeList ((algebra.Alg.Var 6)::nil))::nil))):: R_xml_0_rule_as_list_10. Definition R_xml_0_rule_as_list_12 := ((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_11. Definition R_xml_0_rule_as_list_13 := ((algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil)), (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNeList ((algebra.Alg.Var 5)::nil)):: (algebra.Alg.Var 6)::nil)))::R_xml_0_rule_as_list_12. Definition R_xml_0_rule_as_list_14 := ((algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 6)::nil)), (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_a__isList ((algebra.Alg.Var 6)::nil))::nil))):: R_xml_0_rule_as_list_13. Definition R_xml_0_rule_as_list_15 := ((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_14. Definition R_xml_0_rule_as_list_16 := ((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_a__U62 ((algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Var 4)::nil))::nil))):: R_xml_0_rule_as_list_15. Definition R_xml_0_rule_as_list_17 := ((algebra.Alg.Term algebra.F.id_a__U62 ((algebra.Alg.Term algebra.F.id_tt nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_16. Definition R_xml_0_rule_as_list_18 := ((algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 4)::nil)), (algebra.Alg.Term algebra.F.id_a__U72 ((algebra.Alg.Term algebra.F.id_a__isNePal ((algebra.Alg.Var 4)::nil))::nil))):: R_xml_0_rule_as_list_17. Definition R_xml_0_rule_as_list_19 := ((algebra.Alg.Term algebra.F.id_a__U72 ((algebra.Alg.Term algebra.F.id_tt 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__and ((algebra.Alg.Term algebra.F.id_tt nil)::(algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_19. Definition R_xml_0_rule_as_list_21 := ((algebra.Alg.Term algebra.F.id_a__isList ((algebra.Alg.Var 4)::nil)), (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Var 4)::nil)):: (algebra.Alg.Var 4)::nil)))::R_xml_0_rule_as_list_20. Definition R_xml_0_rule_as_list_22 := ((algebra.Alg.Term algebra.F.id_a__isList ((algebra.Alg.Term algebra.F.id_nil 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__isList ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind ((algebra.Alg.Var 6)::nil))::nil)):: (algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))):: R_xml_0_rule_as_list_22. Definition R_xml_0_rule_as_list_24 := ((algebra.Alg.Term algebra.F.id_a__isNeList ((algebra.Alg.Var 4)::nil)), (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Var 4)::nil)):: (algebra.Alg.Var 4)::nil)))::R_xml_0_rule_as_list_23. Definition R_xml_0_rule_as_list_25 := ((algebra.Alg.Term algebra.F.id_a__isNeList ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::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__isPalListKind ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind ((algebra.Alg.Var 6)::nil))::nil)):: (algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))):: R_xml_0_rule_as_list_24. Definition R_xml_0_rule_as_list_26 := ((algebra.Alg.Term algebra.F.id_a__isNeList ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::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__isPalListKind ((algebra.Alg.Var 5)::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind ((algebra.Alg.Var 6)::nil))::nil)):: (algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))):: R_xml_0_rule_as_list_25. Definition R_xml_0_rule_as_list_27 := ((algebra.Alg.Term algebra.F.id_a__isNePal ((algebra.Alg.Var 4)::nil)), (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Var 4)::nil)):: (algebra.Alg.Var 4)::nil)))::R_xml_0_rule_as_list_26. Definition R_xml_0_rule_as_list_28 := ((algebra.Alg.Term algebra.F.id_a__isNePal ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 7)::(algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 8)::(algebra.Alg.Var 7)::nil))::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind ((algebra.Alg.Var 7)::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isPal ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind ((algebra.Alg.Var 8)::nil))::nil))::nil))):: R_xml_0_rule_as_list_27. Definition R_xml_0_rule_as_list_29 := ((algebra.Alg.Term algebra.F.id_a__isPal ((algebra.Alg.Var 4)::nil)), (algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Var 4)::nil)):: (algebra.Alg.Var 4)::nil)))::R_xml_0_rule_as_list_28. Definition R_xml_0_rule_as_list_30 := ((algebra.Alg.Term algebra.F.id_a__isPal ((algebra.Alg.Term algebra.F.id_nil nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_29. Definition R_xml_0_rule_as_list_31 := ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Term algebra.F.id_a nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_30. Definition R_xml_0_rule_as_list_32 := ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Term algebra.F.id_e nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_31. Definition R_xml_0_rule_as_list_33 := ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Term algebra.F.id_i nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_32. Definition R_xml_0_rule_as_list_34 := ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Term algebra.F.id_nil nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_33. Definition R_xml_0_rule_as_list_35 := ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Term algebra.F.id_o nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_34. Definition R_xml_0_rule_as_list_36 := ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Term algebra.F.id_u nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_35. Definition R_xml_0_rule_as_list_37 := ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Var 5)::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind ((algebra.Alg.Var 6)::nil))::nil)))::R_xml_0_rule_as_list_36. Definition R_xml_0_rule_as_list_38 := ((algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Term algebra.F.id_a nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_37. Definition R_xml_0_rule_as_list_39 := ((algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Term algebra.F.id_e nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_38. Definition R_xml_0_rule_as_list_40 := ((algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Term algebra.F.id_i nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_39. Definition R_xml_0_rule_as_list_41 := ((algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Term algebra.F.id_o nil)::nil)),(algebra.Alg.Term algebra.F.id_tt nil)):: R_xml_0_rule_as_list_40. Definition R_xml_0_rule_as_list_42 := ((algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Term algebra.F.id_u nil)::nil)),(algebra.Alg.Term algebra.F.id_tt 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___ ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 10)::nil))::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_U11 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::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_U12 ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))::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_isNeList ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__isNeList ((algebra.Alg.Var 1)::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_U21 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10):: (algebra.Alg.Var 11)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Var 10):: (algebra.Alg.Var 11)::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_U22 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::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_isList ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__isList ((algebra.Alg.Var 1)::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_U23 ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U23 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))::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_U31 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::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_U32 ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))::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_isQid ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Var 1)::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_U41 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10):: (algebra.Alg.Var 11)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Var 10):: (algebra.Alg.Var 11)::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_U42 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::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_U43 ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))::nil))):: R_xml_0_rule_as_list_55. Definition R_xml_0_rule_as_list_57 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10):: (algebra.Alg.Var 11)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil))::(algebra.Alg.Var 10):: (algebra.Alg.Var 11)::nil)))::R_xml_0_rule_as_list_56. Definition R_xml_0_rule_as_list_58 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_57. Definition R_xml_0_rule_as_list_59 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U53 ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))::nil))):: R_xml_0_rule_as_list_58. Definition R_xml_0_rule_as_list_60 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_59. Definition R_xml_0_rule_as_list_61 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U62 ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U62 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))::nil))):: R_xml_0_rule_as_list_60. Definition R_xml_0_rule_as_list_62 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_61. Definition R_xml_0_rule_as_list_63 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_U72 ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__U72 ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 1)::nil))::nil))):: R_xml_0_rule_as_list_62. Definition R_xml_0_rule_as_list_64 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isNePal ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__isNePal ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_63. Definition R_xml_0_rule_as_list_65 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Var 9)::(algebra.Alg.Var 10)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Var 9)::nil)):: (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_64. Definition R_xml_0_rule_as_list_66 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isPalListKind ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Var 1)::nil)))::R_xml_0_rule_as_list_65. Definition R_xml_0_rule_as_list_67 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_isPal ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_a__isPal ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_66. Definition R_xml_0_rule_as_list_68 := ((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_67. Definition R_xml_0_rule_as_list_69 := ((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_68. Definition R_xml_0_rule_as_list_70 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_a nil)::nil)),(algebra.Alg.Term algebra.F.id_a nil)):: R_xml_0_rule_as_list_69. Definition R_xml_0_rule_as_list_71 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_e nil)::nil)),(algebra.Alg.Term algebra.F.id_e nil)):: R_xml_0_rule_as_list_70. Definition R_xml_0_rule_as_list_72 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_i nil)::nil)),(algebra.Alg.Term algebra.F.id_i nil)):: R_xml_0_rule_as_list_71. Definition R_xml_0_rule_as_list_73 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_o nil)::nil)),(algebra.Alg.Term algebra.F.id_o nil)):: R_xml_0_rule_as_list_72. Definition R_xml_0_rule_as_list_74 := ((algebra.Alg.Term algebra.F.id_mark ((algebra.Alg.Term algebra.F.id_u nil)::nil)),(algebra.Alg.Term algebra.F.id_u nil)):: R_xml_0_rule_as_list_73. Definition R_xml_0_rule_as_list_75 := ((algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)), (algebra.Alg.Term algebra.F.id___ ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_74. Definition R_xml_0_rule_as_list_76 := ((algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)), (algebra.Alg.Term algebra.F.id_U11 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_75. Definition R_xml_0_rule_as_list_77 := ((algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_U12 ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_76. Definition R_xml_0_rule_as_list_78 := ((algebra.Alg.Term algebra.F.id_a__isNeList ((algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_isNeList ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_77. Definition R_xml_0_rule_as_list_79 := ((algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::(algebra.Alg.Var 11)::nil)), (algebra.Alg.Term algebra.F.id_U21 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::(algebra.Alg.Var 11)::nil))):: R_xml_0_rule_as_list_78. Definition R_xml_0_rule_as_list_80 := ((algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)), (algebra.Alg.Term algebra.F.id_U22 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_79. Definition R_xml_0_rule_as_list_81 := ((algebra.Alg.Term algebra.F.id_a__isList ((algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_isList ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_80. Definition R_xml_0_rule_as_list_82 := ((algebra.Alg.Term algebra.F.id_a__U23 ((algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_U23 ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_81. Definition R_xml_0_rule_as_list_83 := ((algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)), (algebra.Alg.Term algebra.F.id_U31 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_82. Definition R_xml_0_rule_as_list_84 := ((algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_U32 ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_83. Definition R_xml_0_rule_as_list_85 := ((algebra.Alg.Term algebra.F.id_a__isQid ((algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_isQid ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_84. Definition R_xml_0_rule_as_list_86 := ((algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::(algebra.Alg.Var 11)::nil)), (algebra.Alg.Term algebra.F.id_U41 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::(algebra.Alg.Var 11)::nil))):: R_xml_0_rule_as_list_85. Definition R_xml_0_rule_as_list_87 := ((algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)), (algebra.Alg.Term algebra.F.id_U42 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_86. Definition R_xml_0_rule_as_list_88 := ((algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_U43 ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_87. Definition R_xml_0_rule_as_list_89 := ((algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::(algebra.Alg.Var 11)::nil)), (algebra.Alg.Term algebra.F.id_U51 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::(algebra.Alg.Var 11)::nil))):: R_xml_0_rule_as_list_88. Definition R_xml_0_rule_as_list_90 := ((algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)), (algebra.Alg.Term algebra.F.id_U52 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_89. Definition R_xml_0_rule_as_list_91 := ((algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_U53 ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_90. Definition R_xml_0_rule_as_list_92 := ((algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)), (algebra.Alg.Term algebra.F.id_U61 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_91. Definition R_xml_0_rule_as_list_93 := ((algebra.Alg.Term algebra.F.id_a__U62 ((algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_U62 ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_92. Definition R_xml_0_rule_as_list_94 := ((algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)), (algebra.Alg.Term algebra.F.id_U71 ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_93. Definition R_xml_0_rule_as_list_95 := ((algebra.Alg.Term algebra.F.id_a__U72 ((algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_U72 ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_94. Definition R_xml_0_rule_as_list_96 := ((algebra.Alg.Term algebra.F.id_a__isNePal ((algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_isNePal ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_95. Definition R_xml_0_rule_as_list_97 := ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)), (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Var 9):: (algebra.Alg.Var 10)::nil)))::R_xml_0_rule_as_list_96. Definition R_xml_0_rule_as_list_98 := ((algebra.Alg.Term algebra.F.id_a__isPalListKind ((algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_isPalListKind ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_97. Definition R_xml_0_rule_as_list_99 := ((algebra.Alg.Term algebra.F.id_a__isPal ((algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id_isPal ((algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_98. Definition R_xml_0_rule_as_list := R_xml_0_rule_as_list_99. 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 100. injection H;intros ;subst;constructor 99. injection H;intros ;subst;constructor 98. injection H;intros ;subst;constructor 97. injection H;intros ;subst;constructor 96. injection H;intros ;subst;constructor 95. injection H;intros ;subst;constructor 94. injection H;intros ;subst;constructor 93. injection H;intros ;subst;constructor 92. injection H;intros ;subst;constructor 91. injection H;intros ;subst;constructor 90. injection H;intros ;subst;constructor 89. injection H;intros ;subst;constructor 88. injection H;intros ;subst;constructor 87. injection H;intros ;subst;constructor 86. injection H;intros ;subst;constructor 85. injection H;intros ;subst;constructor 84. injection H;intros ;subst;constructor 83. injection H;intros ;subst;constructor 82. injection H;intros ;subst;constructor 81. injection H;intros ;subst;constructor 80. injection H;intros ;subst;constructor 79. injection H;intros ;subst;constructor 78. injection H;intros ;subst;constructor 77. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. 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. injection H;intros ;subst;constructor 54. injection H;intros ;subst;constructor 53. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. 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. injection H;intros ;subst;constructor 30. injection H;intros ;subst;constructor 29. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. 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. injection H;intros ;subst;constructor 6. injection H;intros ;subst;constructor 5. rewrite <- or_ext_generated.or5_equiv in H|-. case H;clear H;intros H. 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_32 (x13 x14 x15 x16 x17 x18 x19 x20 x21 x22 x23 x24 x25 x26 x27 x28 x29 x30 x31 x32 x33 x34 x35 x36 x37 x38 x39 x40 x41 x42 x43 x44:Prop) : Prop := | conj_32 : x13->x14->x15->x16->x17->x18->x19->x20->x21->x22->x23->x24->x25-> x26->x27->x28->x29->x30->x31->x32->x33->x34->x35->x36->x37->x38-> x39->x40->x41->x42->x43->x44-> and_32 x13 x14 x15 x16 x17 x18 x19 x20 x21 x22 x23 x24 x25 x26 x27 x28 x29 x30 x31 x32 x33 x34 x35 x36 x37 x38 x39 x40 x41 x42 x43 x44 . 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_32 (t = (algebra.Alg.Term algebra.F.id_a nil) -> t' = (algebra.Alg.Term algebra.F.id_a nil)) (forall x14 x16, t = (algebra.Alg.Term algebra.F.id_U42 (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U42 (x13::x15::nil))/\ (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 x14 x16 x18, t = (algebra.Alg.Term algebra.F.id_U21 (x14::x16::x18::nil)) -> exists x13, exists x15, exists x17, t' = (algebra.Alg.Term algebra.F.id_U21 (x13::x15::x17::nil))/\ (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)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x17 x18)) (forall x14, t = (algebra.Alg.Term algebra.F.id_U72 (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_U72 (x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (t = (algebra.Alg.Term algebra.F.id_u nil) -> t' = (algebra.Alg.Term algebra.F.id_u nil)) (forall x14, t = (algebra.Alg.Term algebra.F.id_U53 (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_U53 (x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x14 x16, t = (algebra.Alg.Term algebra.F.id_U31 (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U31 (x13::x15::nil))/\ (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 x14, t = (algebra.Alg.Term algebra.F.id_isPalListKind (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (t = (algebra.Alg.Term algebra.F.id_i nil) -> t' = (algebra.Alg.Term algebra.F.id_i nil)) (forall x14 x16 x18, t = (algebra.Alg.Term algebra.F.id_U51 (x14::x16::x18::nil)) -> exists x13, exists x15, exists x17, t' = (algebra.Alg.Term algebra.F.id_U51 (x13::x15::x17::nil))/\ (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)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x17 x18)) (forall x14, t = (algebra.Alg.Term algebra.F.id_isList (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_isList (x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x14, t = (algebra.Alg.Term algebra.F.id_U12 (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_U12 (x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x14, t = (algebra.Alg.Term algebra.F.id_U62 (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_U62 (x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x14, t = (algebra.Alg.Term algebra.F.id_isQid (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_isQid (x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x14, t = (algebra.Alg.Term algebra.F.id_isPal (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_isPal (x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x14 x16, t = (algebra.Alg.Term algebra.F.id___ (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id___ (x13::x15::nil))/\ (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)) (t = (algebra.Alg.Term algebra.F.id_e nil) -> t' = (algebra.Alg.Term algebra.F.id_e nil)) (forall x14, t = (algebra.Alg.Term algebra.F.id_U43 (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_U43 (x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x14 x16, t = (algebra.Alg.Term algebra.F.id_U22 (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U22 (x13::x15::nil))/\ (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 x14, t = (algebra.Alg.Term algebra.F.id_isNePal (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_isNePal (x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (t = (algebra.Alg.Term algebra.F.id_tt nil) -> t' = (algebra.Alg.Term algebra.F.id_tt nil)) (forall x14 x16, t = (algebra.Alg.Term algebra.F.id_U11 (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U11 (x13::x15::nil))/\ (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 x14 x16, t = (algebra.Alg.Term algebra.F.id_U61 (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U61 (x13::x15::nil))/\ (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 x14, t = (algebra.Alg.Term algebra.F.id_U32 (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_U32 (x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x14 x16, t = (algebra.Alg.Term algebra.F.id_and (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_and (x13::x15::nil))/\ (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)) (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_o nil) -> t' = (algebra.Alg.Term algebra.F.id_o nil)) (forall x14 x16, t = (algebra.Alg.Term algebra.F.id_U52 (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U52 (x13::x15::nil))/\ (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 x14, t = (algebra.Alg.Term algebra.F.id_U23 (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_U23 (x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x14, t = (algebra.Alg.Term algebra.F.id_isNeList (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_isNeList (x13::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x13 x14)) (forall x14 x16, t = (algebra.Alg.Term algebra.F.id_U71 (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U71 (x13::x15::nil))/\ (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 x14 x16 x18, t = (algebra.Alg.Term algebra.F.id_U41 (x14::x16::x18::nil)) -> exists x13, exists x15, exists x17, t' = (algebra.Alg.Term algebra.F.id_U41 (x13::x15::x17::nil))/\ (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)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x17 x18)). Proof. intros t t' H. induction H as [|y IH z z_to_y] using closure_extension.refl_trans_clos_ind2. constructor 1. intros H;intuition;constructor 1. intros x14 x16 H;exists x14;exists x16;intuition;constructor 1. intros x14 x16 x18 H;exists x14;exists x16;exists x18;intuition; constructor 1. intros x14 H;exists x14;intuition;constructor 1. intros H;intuition;constructor 1. intros x14 H;exists x14;intuition;constructor 1. intros x14 x16 H;exists x14;exists x16;intuition;constructor 1. intros x14 H;exists x14;intuition;constructor 1. intros H;intuition;constructor 1. intros x14 x16 x18 H;exists x14;exists x16;exists x18;intuition; constructor 1. intros x14 H;exists x14;intuition;constructor 1. intros x14 H;exists x14;intuition;constructor 1. intros x14 H;exists x14;intuition;constructor 1. intros x14 H;exists x14;intuition;constructor 1. intros x14 H;exists x14;intuition;constructor 1. intros x14 x16 H;exists x14;exists x16;intuition;constructor 1. intros H;intuition;constructor 1. intros x14 H;exists x14;intuition;constructor 1. intros x14 x16 H;exists x14;exists x16;intuition;constructor 1. intros x14 H;exists x14;intuition;constructor 1. intros H;intuition;constructor 1. intros x14 x16 H;exists x14;exists x16;intuition;constructor 1. intros x14 x16 H;exists x14;exists x16;intuition;constructor 1. intros x14 H;exists x14;intuition;constructor 1. intros x14 x16 H;exists x14;exists x16;intuition;constructor 1. intros H;intuition;constructor 1. intros H;intuition;constructor 1. intros x14 x16 H;exists x14;exists x16;intuition;constructor 1. intros x14 H;exists x14;intuition;constructor 1. intros x14 H;exists x14;intuition;constructor 1. intros x14 x16 H;exists x14;exists x16;intuition;constructor 1. intros x14 x16 x18 H;exists x14;exists x16;exists x18;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_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41]. constructor. clear H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;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_a H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_U42 y x16 (refl_equal _)) as [x13 [x15]];intros ; intuition;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_U42 x14 y (refl_equal _)) as [x13 [x15]];intros ; intuition;exists x13;exists x15;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_U42 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros x14 x16 x18 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 x14 |- _ => destruct (H_id_U21 y x16 x18 (refl_equal _)) as [x13 [x15 [x17]]]; intros ;intuition;exists x13;exists x15;exists x17;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x16 |- _ => destruct (H_id_U21 x14 y x18 (refl_equal _)) as [x13 [x15 [x17]]]; intros ;intuition;exists x13;exists x15;exists x17;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x18 |- _ => destruct (H_id_U21 x14 x16 y (refl_equal _)) as [x13 [x15 [x17]]]; intros ;intuition;exists x13;exists x15;exists x17;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_U72 y (refl_equal _)) as [x13];intros ;intuition; exists x13;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;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_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_U53 y (refl_equal _)) as [x13];intros ;intuition; exists x13;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_U31 y x16 (refl_equal _)) as [x13 [x15]];intros ; intuition;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_U31 x14 y (refl_equal _)) as [x13 [x15]];intros ; intuition;exists x13;exists x15;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_isPalListKind y (refl_equal _)) as [x13];intros ; intuition;exists x13;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;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_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros x14 x16 x18 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 x14 |- _ => destruct (H_id_U51 y x16 x18 (refl_equal _)) as [x13 [x15 [x17]]]; intros ;intuition;exists x13;exists x15;exists x17;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x16 |- _ => destruct (H_id_U51 x14 y x18 (refl_equal _)) as [x13 [x15 [x17]]]; intros ;intuition;exists x13;exists x15;exists x17;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x18 |- _ => destruct (H_id_U51 x14 x16 y (refl_equal _)) as [x13 [x15 [x17]]]; intros ;intuition;exists x13;exists x15;exists x17;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_isList y (refl_equal _)) as [x13];intros ;intuition; exists x13;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_U12 y (refl_equal _)) as [x13];intros ;intuition; exists x13;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_U62 y (refl_equal _)) as [x13];intros ;intuition; exists x13;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_isQid y (refl_equal _)) as [x13];intros ;intuition; exists x13;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_isPal y (refl_equal _)) as [x13];intros ;intuition; exists x13;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id___ y x16 (refl_equal _)) as [x13 [x15]];intros ; intuition;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___ x14 y (refl_equal _)) as [x13 [x15]];intros ; intuition;exists x13;exists x15;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;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_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_U43 y (refl_equal _)) as [x13];intros ;intuition; exists x13;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_U22 y x16 (refl_equal _)) as [x13 [x15]];intros ; intuition;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_U22 x14 y (refl_equal _)) as [x13 [x15]];intros ; intuition;exists x13;exists x15;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_isNePal y (refl_equal _)) as [x13];intros ;intuition; exists x13;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;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_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_U11 y x16 (refl_equal _)) as [x13 [x15]];intros ; intuition;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_U11 x14 y (refl_equal _)) as [x13 [x15]];intros ; intuition;exists x13;exists x15;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_U61 y x16 (refl_equal _)) as [x13 [x15]];intros ; intuition;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_U61 x14 y (refl_equal _)) as [x13 [x15]];intros ; intuition;exists x13;exists x15;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_U32 y (refl_equal _)) as [x13];intros ;intuition; exists x13;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_and y x16 (refl_equal _)) as [x13 [x15]];intros ; intuition;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_and x14 y (refl_equal _)) as [x13 [x15]];intros ; intuition;exists x13;exists x15;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;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_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_U52 H_id_U23 H_id_isNeList H_id_U71 H_id_U41;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_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U23 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_U52 y x16 (refl_equal _)) as [x13 [x15]];intros ; intuition;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_U52 x14 y (refl_equal _)) as [x13 [x15]];intros ; intuition;exists x13;exists x15;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_isNeList H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_U23 y (refl_equal _)) as [x13];intros ;intuition; exists x13;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_U71 H_id_U41;intros 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 x14 |- _ => destruct (H_id_isNeList y (refl_equal _)) as [x13];intros ;intuition; exists x13;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U41;intros 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 x14 |- _ => destruct (H_id_U71 y x16 (refl_equal _)) as [x13 [x15]];intros ; intuition;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_U71 x14 y (refl_equal _)) as [x13 [x15]];intros ; intuition;exists x13;exists x15;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_U42 H_id_U21 H_id_U72 H_id_u H_id_U53 H_id_U31 H_id_isPalListKind H_id_i H_id_U51 H_id_isList H_id_U12 H_id_U62 H_id_isQid H_id_isPal H_id___ H_id_e H_id_U43 H_id_U22 H_id_isNePal H_id_tt H_id_U11 H_id_U61 H_id_U32 H_id_and H_id_nil H_id_o H_id_U52 H_id_U23 H_id_isNeList H_id_U71;intros x14 x16 x18 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 x14 |- _ => destruct (H_id_U41 y x16 x18 (refl_equal _)) as [x13 [x15 [x17]]]; intros ;intuition;exists x13;exists x15;exists x17;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x16 |- _ => destruct (H_id_U41 x14 y x18 (refl_equal _)) as [x13 [x15 [x17]]]; intros ;intuition;exists x13;exists x15;exists x17;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x18 |- _ => destruct (H_id_U41 x14 x16 y (refl_equal _)) as [x13 [x15 [x17]]]; intros ;intuition;exists x13;exists x15;exists x17;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . Qed. Lemma id_a_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> t = (algebra.Alg.Term algebra.F.id_a nil) -> t' = (algebra.Alg.Term algebra.F.id_a nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U42_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14 x16, t = (algebra.Alg.Term algebra.F.id_U42 (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U42 (x13::x15::nil))/\ (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 x14 x16 x18, t = (algebra.Alg.Term algebra.F.id_U21 (x14::x16::x18::nil)) -> exists x13, exists x15, exists x17, t' = (algebra.Alg.Term algebra.F.id_U21 (x13::x15::x17::nil))/\ (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)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x17 x18). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U72_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14, t = (algebra.Alg.Term algebra.F.id_U72 (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_U72 (x13::nil))/\ (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_u_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> t = (algebra.Alg.Term algebra.F.id_u nil) -> t' = (algebra.Alg.Term algebra.F.id_u nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U53_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14, t = (algebra.Alg.Term algebra.F.id_U53 (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_U53 (x13::nil))/\ (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_U31_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14 x16, t = (algebra.Alg.Term algebra.F.id_U31 (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U31 (x13::x15::nil))/\ (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_isPalListKind_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14, t = (algebra.Alg.Term algebra.F.id_isPalListKind (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil))/\ (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_i_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> t = (algebra.Alg.Term algebra.F.id_i nil) -> t' = (algebra.Alg.Term algebra.F.id_i nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U51_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14 x16 x18, t = (algebra.Alg.Term algebra.F.id_U51 (x14::x16::x18::nil)) -> exists x13, exists x15, exists x17, t' = (algebra.Alg.Term algebra.F.id_U51 (x13::x15::x17::nil))/\ (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)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x17 x18). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_isList_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14, t = (algebra.Alg.Term algebra.F.id_isList (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_isList (x13::nil))/\ (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_U12_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14, t = (algebra.Alg.Term algebra.F.id_U12 (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_U12 (x13::nil))/\ (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_U62_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14, t = (algebra.Alg.Term algebra.F.id_U62 (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_U62 (x13::nil))/\ (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_isQid_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14, t = (algebra.Alg.Term algebra.F.id_isQid (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_isQid (x13::nil))/\ (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_isPal_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14, t = (algebra.Alg.Term algebra.F.id_isPal (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_isPal (x13::nil))/\ (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____is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14 x16, t = (algebra.Alg.Term algebra.F.id___ (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id___ (x13::x15::nil))/\ (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_e_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> t = (algebra.Alg.Term algebra.F.id_e nil) -> t' = (algebra.Alg.Term algebra.F.id_e nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. 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 x14, t = (algebra.Alg.Term algebra.F.id_U43 (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_U43 (x13::nil))/\ (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_U22_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14 x16, t = (algebra.Alg.Term algebra.F.id_U22 (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U22 (x13::x15::nil))/\ (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_isNePal_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14, t = (algebra.Alg.Term algebra.F.id_isNePal (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_isNePal (x13::nil))/\ (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_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_U11_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14 x16, t = (algebra.Alg.Term algebra.F.id_U11 (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U11 (x13::x15::nil))/\ (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_U61_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14 x16, t = (algebra.Alg.Term algebra.F.id_U61 (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U61 (x13::x15::nil))/\ (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_U32_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14, t = (algebra.Alg.Term algebra.F.id_U32 (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_U32 (x13::nil))/\ (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_and_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14 x16, t = (algebra.Alg.Term algebra.F.id_and (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_and (x13::x15::nil))/\ (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_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_o_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> t = (algebra.Alg.Term algebra.F.id_o nil) -> t' = (algebra.Alg.Term algebra.F.id_o nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_U52_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14 x16, t = (algebra.Alg.Term algebra.F.id_U52 (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U52 (x13::x15::nil))/\ (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_U23_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14, t = (algebra.Alg.Term algebra.F.id_U23 (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_U23 (x13::nil))/\ (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_isNeList_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14, t = (algebra.Alg.Term algebra.F.id_isNeList (x14::nil)) -> exists x13, t' = (algebra.Alg.Term algebra.F.id_isNeList (x13::nil))/\ (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_U71_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14 x16, t = (algebra.Alg.Term algebra.F.id_U71 (x14::x16::nil)) -> exists x13, exists x15, t' = (algebra.Alg.Term algebra.F.id_U71 (x13::x15::nil))/\ (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_U41_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x14 x16 x18, t = (algebra.Alg.Term algebra.F.id_U41 (x14::x16::x18::nil)) -> exists x13, exists x15, exists x17, t' = (algebra.Alg.Term algebra.F.id_U41 (x13::x15::x17::nil))/\ (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)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x17 x18). 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_a nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_a_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; impossible_star_reduction_R_xml_0 )) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U42 (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := 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 [x14 [x13 [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_U21 (?x15::?x14::?x13::nil)) |- _ => let x15 := fresh "x" in (let x14 := fresh "x" in (let x13 := 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_U21_is_R_xml_0_constructor H (refl_equal _)) as [x15 [x14 [x13 [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_U72 (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_U72_is_R_xml_0_constructor H (refl_equal _)) as [x13 [Heq Hred1]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_u nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_u_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; impossible_star_reduction_R_xml_0 )) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U53 (?x13::nil)) |- _ => let x13 := 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 [x13 [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 (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := 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 [x14 [x13 [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_isPalListKind (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isPalListKind_is_R_xml_0_constructor H (refl_equal _)) as [x13 [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_i nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_i_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; impossible_star_reduction_R_xml_0 )) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U51 (?x15::?x14::?x13::nil)) |- _ => let x15 := fresh "x" in (let x14 := fresh "x" in (let x13 := 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 [x15 [x14 [x13 [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_isList (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isList_is_R_xml_0_constructor H (refl_equal _)) as [x13 [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_U12 (?x13::nil)) |- _ => let x13 := 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 [x13 [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_U62 (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_U62_is_R_xml_0_constructor H (refl_equal _)) as [x13 [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_isQid (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isQid_is_R_xml_0_constructor H (refl_equal _)) as [x13 [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_isPal (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isPal_is_R_xml_0_constructor H (refl_equal _)) as [x13 [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___ (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id____is_R_xml_0_constructor H (refl_equal _)) as [x14 [x13 [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_e nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_e_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; impossible_star_reduction_R_xml_0 )) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U43 (?x13::nil)) |- _ => let x13 := 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 [x13 [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_U22 (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_U22_is_R_xml_0_constructor H (refl_equal _)) as [x14 [x13 [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_isNePal (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isNePal_is_R_xml_0_constructor H (refl_equal _)) as [x13 [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_U11 (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := 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 [x14 [x13 [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_U61 (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := 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 [x14 [x13 [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_U32 (?x13::nil)) |- _ => let x13 := 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 [x13 [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 (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := 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 [x14 [x13 [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_o nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_o_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; impossible_star_reduction_R_xml_0 )) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U52 (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := 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 [x14 [x13 [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_U23 (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_U23_is_R_xml_0_constructor H (refl_equal _)) as [x13 [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_isNeList (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isNeList_is_R_xml_0_constructor H (refl_equal _)) as [x13 [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_U71 (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_U71_is_R_xml_0_constructor H (refl_equal _)) as [x14 [x13 [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_U41 (?x15::?x14::?x13::nil)) |- _ => let x15 := fresh "x" in (let x14 := fresh "x" in (let x13 := 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 [x15 [x14 [x13 [Heq [Hred3 [Hred2 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_a nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_a_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; try (simplify_star_reduction_R_xml_0 ))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U42 (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := 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 [x14 [x13 [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_U21 (?x15::?x14::?x13::nil)) |- _ => let x15 := fresh "x" in (let x14 := fresh "x" in (let x13 := 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_U21_is_R_xml_0_constructor H (refl_equal _)) as [x15 [x14 [x13 [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_U72 (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_U72_is_R_xml_0_constructor H (refl_equal _)) as [x13 [Heq Hred1]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_u nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_u_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; try (simplify_star_reduction_R_xml_0 ))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U53 (?x13::nil)) |- _ => let x13 := 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 [x13 [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 (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := 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 [x14 [x13 [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_isPalListKind (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isPalListKind_is_R_xml_0_constructor H (refl_equal _)) as [x13 [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_i nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_i_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; try (simplify_star_reduction_R_xml_0 ))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U51 (?x15::?x14::?x13::nil)) |- _ => let x15 := fresh "x" in (let x14 := fresh "x" in (let x13 := 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 [x15 [x14 [x13 [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_isList (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isList_is_R_xml_0_constructor H (refl_equal _)) as [x13 [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_U12 (?x13::nil)) |- _ => let x13 := 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 [x13 [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_U62 (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_U62_is_R_xml_0_constructor H (refl_equal _)) as [x13 [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_isQid (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isQid_is_R_xml_0_constructor H (refl_equal _)) as [x13 [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_isPal (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isPal_is_R_xml_0_constructor H (refl_equal _)) as [x13 [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___ (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id____is_R_xml_0_constructor H (refl_equal _)) as [x14 [x13 [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_e nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_e_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; try (simplify_star_reduction_R_xml_0 ))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U43 (?x13::nil)) |- _ => let x13 := 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 [x13 [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_U22 (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_U22_is_R_xml_0_constructor H (refl_equal _)) as [x14 [x13 [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_isNePal (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isNePal_is_R_xml_0_constructor H (refl_equal _)) as [x13 [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_U11 (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := 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 [x14 [x13 [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_U61 (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := 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 [x14 [x13 [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_U32 (?x13::nil)) |- _ => let x13 := 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 [x13 [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 (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := 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 [x14 [x13 [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_o nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_o_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; try (simplify_star_reduction_R_xml_0 ))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_U52 (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := 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 [x14 [x13 [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_U23 (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_U23_is_R_xml_0_constructor H (refl_equal _)) as [x13 [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_isNeList (?x13::nil)) |- _ => let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_isNeList_is_R_xml_0_constructor H (refl_equal _)) as [x13 [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_U71 (?x14::?x13::nil)) |- _ => let x14 := fresh "x" in (let x13 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_U71_is_R_xml_0_constructor H (refl_equal _)) as [x14 [x13 [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_U41 (?x15::?x14::?x13::nil)) |- _ => let x15 := fresh "x" in (let x14 := fresh "x" in (let x13 := 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 [x15 [x14 [x13 [Heq [Hred3 [Hred2 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____ : A ->A ->A. Hypothesis P_id_a : A. Hypothesis P_id_a__U42 : A ->A ->A. Hypothesis P_id_U42 : A ->A ->A. Hypothesis P_id_a__U21 : A ->A ->A ->A. Hypothesis P_id_U21 : A ->A ->A ->A. Hypothesis P_id_a__U72 : A ->A. Hypothesis P_id_U72 : A ->A. Hypothesis P_id_a__U11 : A ->A ->A. Hypothesis P_id_u : A. Hypothesis P_id_a__U53 : A ->A. Hypothesis P_id_U53 : A ->A. Hypothesis P_id_a__U31 : A ->A ->A. Hypothesis P_id_U31 : A ->A ->A. Hypothesis P_id_isPalListKind : A ->A. Hypothesis P_id_mark : A ->A. Hypothesis P_id_i : A. Hypothesis P_id_a__U51 : A ->A ->A ->A. Hypothesis P_id_U51 : A ->A ->A ->A. Hypothesis P_id_a__isList : A ->A. Hypothesis P_id_isList : A ->A. Hypothesis P_id_a__and : A ->A ->A. Hypothesis P_id_a__U12 : A ->A. Hypothesis P_id_U12 : A ->A. Hypothesis P_id_a__U62 : A ->A. Hypothesis P_id_U62 : A ->A. Hypothesis P_id_a__isQid : A ->A. Hypothesis P_id_isQid : A ->A. Hypothesis P_id_isPal : A ->A. Hypothesis P_id___ : A ->A ->A. Hypothesis P_id_e : A. Hypothesis P_id_a__U43 : A ->A. Hypothesis P_id_U43 : A ->A. Hypothesis P_id_a__U22 : A ->A ->A. Hypothesis P_id_U22 : A ->A ->A. Hypothesis P_id_a__isNePal : A ->A. Hypothesis P_id_isNePal : A ->A. Hypothesis P_id_tt : A. Hypothesis P_id_U11 : A ->A ->A. Hypothesis P_id_a__U61 : A ->A ->A. Hypothesis P_id_U61 : A ->A ->A. Hypothesis P_id_a__U32 : A ->A. Hypothesis P_id_U32 : A ->A. Hypothesis P_id_and : A ->A ->A. Hypothesis P_id_nil : A. Hypothesis P_id_o : A. Hypothesis P_id_a__U52 : A ->A ->A. Hypothesis P_id_U52 : A ->A ->A. Hypothesis P_id_a__U23 : A ->A. Hypothesis P_id_U23 : A ->A. Hypothesis P_id_a__isPalListKind : A ->A. Hypothesis P_id_a__isNeList : A ->A. Hypothesis P_id_isNeList : A ->A. Hypothesis P_id_a__U71 : A ->A ->A. Hypothesis P_id_U71 : A ->A ->A. Hypothesis P_id_a__U41 : A ->A ->A ->A. Hypothesis P_id_U41 : A ->A ->A ->A. Hypothesis P_id_a__isPal : A ->A. Hypothesis P_id_a_____monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_a____ x14 x16 <= P_id_a____ x13 x15. Hypothesis P_id_a__U42_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_a__U42 x14 x16 <= P_id_a__U42 x13 x15. Hypothesis P_id_U42_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_U42 x14 x16 <= P_id_U42 x13 x15. Hypothesis P_id_a__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) -> P_id_a__U21 x14 x16 x18 <= P_id_a__U21 x13 x15 x17. Hypothesis P_id_U21_monotonic : forall x16 x18 x14 x17 x13 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) -> P_id_U21 x14 x16 x18 <= P_id_U21 x13 x15 x17. Hypothesis P_id_a__U72_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_a__U72 x14 <= P_id_a__U72 x13. Hypothesis P_id_U72_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_U72 x14 <= P_id_U72 x13. Hypothesis P_id_a__U11_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_a__U11 x14 x16 <= P_id_a__U11 x13 x15. Hypothesis P_id_a__U53_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_a__U53 x14 <= P_id_a__U53 x13. Hypothesis P_id_U53_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_U53 x14 <= P_id_U53 x13. Hypothesis P_id_a__U31_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_a__U31 x14 x16 <= P_id_a__U31 x13 x15. Hypothesis P_id_U31_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_U31 x14 x16 <= P_id_U31 x13 x15. Hypothesis P_id_isPalListKind_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) -> P_id_isPalListKind x14 <= P_id_isPalListKind x13. Hypothesis P_id_mark_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_mark x14 <= P_id_mark x13. Hypothesis P_id_a__U51_monotonic : forall x16 x18 x14 x17 x13 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) -> P_id_a__U51 x14 x16 x18 <= P_id_a__U51 x13 x15 x17. Hypothesis P_id_U51_monotonic : forall x16 x18 x14 x17 x13 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) -> P_id_U51 x14 x16 x18 <= P_id_U51 x13 x15 x17. Hypothesis P_id_a__isList_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_a__isList x14 <= P_id_a__isList x13. Hypothesis P_id_isList_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_isList x14 <= P_id_isList x13. Hypothesis P_id_a__and_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_a__and x14 x16 <= P_id_a__and x13 x15. Hypothesis P_id_a__U12_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_a__U12 x14 <= P_id_a__U12 x13. Hypothesis P_id_U12_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_U12 x14 <= P_id_U12 x13. Hypothesis P_id_a__U62_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_a__U62 x14 <= P_id_a__U62 x13. Hypothesis P_id_U62_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_U62 x14 <= P_id_U62 x13. Hypothesis P_id_a__isQid_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_a__isQid x14 <= P_id_a__isQid x13. Hypothesis P_id_isQid_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_isQid x14 <= P_id_isQid x13. Hypothesis P_id_isPal_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_isPal x14 <= P_id_isPal x13. Hypothesis P_id____monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id___ x14 x16 <= P_id___ x13 x15. Hypothesis P_id_a__U43_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_a__U43 x14 <= P_id_a__U43 x13. Hypothesis P_id_U43_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_U43 x14 <= P_id_U43 x13. Hypothesis P_id_a__U22_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_a__U22 x14 x16 <= P_id_a__U22 x13 x15. Hypothesis P_id_U22_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_U22 x14 x16 <= P_id_U22 x13 x15. Hypothesis P_id_a__isNePal_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_a__isNePal x14 <= P_id_a__isNePal x13. Hypothesis P_id_isNePal_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_isNePal x14 <= P_id_isNePal x13. Hypothesis P_id_U11_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_U11 x14 x16 <= P_id_U11 x13 x15. Hypothesis P_id_a__U61_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_a__U61 x14 x16 <= P_id_a__U61 x13 x15. Hypothesis P_id_U61_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_U61 x14 x16 <= P_id_U61 x13 x15. Hypothesis P_id_a__U32_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_a__U32 x14 <= P_id_a__U32 x13. Hypothesis P_id_U32_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_U32 x14 <= P_id_U32 x13. Hypothesis P_id_and_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_and x14 x16 <= P_id_and x13 x15. Hypothesis P_id_a__U52_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_a__U52 x14 x16 <= P_id_a__U52 x13 x15. Hypothesis P_id_U52_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_U52 x14 x16 <= P_id_U52 x13 x15. Hypothesis P_id_a__U23_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_a__U23 x14 <= P_id_a__U23 x13. Hypothesis P_id_U23_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_U23 x14 <= P_id_U23 x13. Hypothesis P_id_a__isPalListKind_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) -> P_id_a__isPalListKind x14 <= P_id_a__isPalListKind x13. Hypothesis P_id_a__isNeList_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) -> P_id_a__isNeList x14 <= P_id_a__isNeList x13. Hypothesis P_id_isNeList_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_isNeList x14 <= P_id_isNeList x13. Hypothesis P_id_a__U71_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_a__U71 x14 x16 <= P_id_a__U71 x13 x15. Hypothesis P_id_U71_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_U71 x14 x16 <= P_id_U71 x13 x15. Hypothesis P_id_a__U41_monotonic : forall x16 x18 x14 x17 x13 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) -> P_id_a__U41 x14 x16 x18 <= P_id_a__U41 x13 x15 x17. Hypothesis P_id_U41_monotonic : forall x16 x18 x14 x17 x13 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) -> P_id_U41 x14 x16 x18 <= P_id_U41 x13 x15 x17. Hypothesis P_id_a__isPal_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_a__isPal x14 <= P_id_a__isPal x13. Hypothesis P_id_a_____bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_a____ x14 x13. Hypothesis P_id_a_bounded : A0 <= P_id_a . Hypothesis P_id_a__U42_bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_a__U42 x14 x13. Hypothesis P_id_U42_bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_U42 x14 x13. Hypothesis P_id_a__U21_bounded : forall x14 x13 x15, (A0 <= x13) ->(A0 <= x14) ->(A0 <= x15) ->A0 <= P_id_a__U21 x15 x14 x13. Hypothesis P_id_U21_bounded : forall x14 x13 x15, (A0 <= x13) ->(A0 <= x14) ->(A0 <= x15) ->A0 <= P_id_U21 x15 x14 x13. Hypothesis P_id_a__U72_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_a__U72 x13. Hypothesis P_id_U72_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_U72 x13. Hypothesis P_id_a__U11_bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_a__U11 x14 x13. Hypothesis P_id_u_bounded : A0 <= P_id_u . Hypothesis P_id_a__U53_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_a__U53 x13. Hypothesis P_id_U53_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_U53 x13. Hypothesis P_id_a__U31_bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_a__U31 x14 x13. Hypothesis P_id_U31_bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_U31 x14 x13. Hypothesis P_id_isPalListKind_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_isPalListKind x13. Hypothesis P_id_mark_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_mark x13. Hypothesis P_id_i_bounded : A0 <= P_id_i . Hypothesis P_id_a__U51_bounded : forall x14 x13 x15, (A0 <= x13) ->(A0 <= x14) ->(A0 <= x15) ->A0 <= P_id_a__U51 x15 x14 x13. Hypothesis P_id_U51_bounded : forall x14 x13 x15, (A0 <= x13) ->(A0 <= x14) ->(A0 <= x15) ->A0 <= P_id_U51 x15 x14 x13. Hypothesis P_id_a__isList_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_a__isList x13. Hypothesis P_id_isList_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_isList x13. Hypothesis P_id_a__and_bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_a__and x14 x13. Hypothesis P_id_a__U12_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_a__U12 x13. Hypothesis P_id_U12_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_U12 x13. Hypothesis P_id_a__U62_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_a__U62 x13. Hypothesis P_id_U62_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_U62 x13. Hypothesis P_id_a__isQid_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_a__isQid x13. Hypothesis P_id_isQid_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_isQid x13. Hypothesis P_id_isPal_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_isPal x13. Hypothesis P_id____bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id___ x14 x13. Hypothesis P_id_e_bounded : A0 <= P_id_e . Hypothesis P_id_a__U43_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_a__U43 x13. Hypothesis P_id_U43_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_U43 x13. Hypothesis P_id_a__U22_bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_a__U22 x14 x13. Hypothesis P_id_U22_bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_U22 x14 x13. Hypothesis P_id_a__isNePal_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_a__isNePal x13. Hypothesis P_id_isNePal_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_isNePal x13. Hypothesis P_id_tt_bounded : A0 <= P_id_tt . Hypothesis P_id_U11_bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_U11 x14 x13. Hypothesis P_id_a__U61_bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_a__U61 x14 x13. Hypothesis P_id_U61_bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_U61 x14 x13. Hypothesis P_id_a__U32_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_a__U32 x13. Hypothesis P_id_U32_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_U32 x13. Hypothesis P_id_and_bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_and x14 x13. Hypothesis P_id_nil_bounded : A0 <= P_id_nil . Hypothesis P_id_o_bounded : A0 <= P_id_o . Hypothesis P_id_a__U52_bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_a__U52 x14 x13. Hypothesis P_id_U52_bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_U52 x14 x13. Hypothesis P_id_a__U23_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_a__U23 x13. Hypothesis P_id_U23_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_U23 x13. Hypothesis P_id_a__isPalListKind_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_a__isPalListKind x13. Hypothesis P_id_a__isNeList_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_a__isNeList x13. Hypothesis P_id_isNeList_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_isNeList x13. Hypothesis P_id_a__U71_bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_a__U71 x14 x13. Hypothesis P_id_U71_bounded : forall x14 x13, (A0 <= x13) ->(A0 <= x14) ->A0 <= P_id_U71 x14 x13. Hypothesis P_id_a__U41_bounded : forall x14 x13 x15, (A0 <= x13) ->(A0 <= x14) ->(A0 <= x15) ->A0 <= P_id_a__U41 x15 x14 x13. Hypothesis P_id_U41_bounded : forall x14 x13 x15, (A0 <= x13) ->(A0 <= x14) ->(A0 <= x15) ->A0 <= P_id_U41 x15 x14 x13. Hypothesis P_id_a__isPal_bounded : forall x13, (A0 <= x13) ->A0 <= P_id_a__isPal x13. Fixpoint measure t { struct t } := match t with | (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) => P_id_a____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) => P_id_a__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U42 (x14::x13::nil)) => P_id_U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) => P_id_a__U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U21 (x15::x14::x13::nil)) => P_id_U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_a__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_U72 (x13::nil)) => P_id_U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) => P_id_a__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_a__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_U53 (x13::nil)) => P_id_U53 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) => P_id_a__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U31 (x14::x13::nil)) => P_id_U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil)) => P_id_isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_mark (measure x13) | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) => P_id_a__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U51 (x15::x14::x13::nil)) => P_id_U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_a__isList (measure x13) | (algebra.Alg.Term algebra.F.id_isList (x13::nil)) => P_id_isList (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) => P_id_a__and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_a__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_U12 (x13::nil)) => P_id_U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_a__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_U62 (x13::nil)) => P_id_U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_a__isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isQid (x13::nil)) => P_id_isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isPal (x13::nil)) => P_id_isPal (measure x13) | (algebra.Alg.Term algebra.F.id___ (x14::x13::nil)) => P_id___ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_a__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_U43 (x13::nil)) => P_id_U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) => P_id_a__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U22 (x14::x13::nil)) => P_id_U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_a__isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_isNePal (x13::nil)) => P_id_isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U11 (x14::x13::nil)) => P_id_U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) => P_id_a__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U61 (x14::x13::nil)) => P_id_U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_a__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_U32 (x13::nil)) => P_id_U32 (measure x13) | (algebra.Alg.Term algebra.F.id_and (x14::x13::nil)) => P_id_and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o | (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) => P_id_a__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U52 (x14::x13::nil)) => P_id_U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_a__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_U23 (x13::nil)) => P_id_U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_a__isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_a__isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_isNeList (x13::nil)) => P_id_isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) => P_id_a__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U71 (x14::x13::nil)) => P_id_U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) => P_id_a__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U41 (x15::x14::x13::nil)) => P_id_U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_a__isPal (measure x13) | _ => A0 end. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) => P_id_a____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) => P_id_a__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U42 (x14::x13::nil)) => P_id_U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14:: x13::nil)) => P_id_a__U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U21 (x15::x14::x13::nil)) => P_id_U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_a__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_U72 (x13::nil)) => P_id_U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) => P_id_a__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_a__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_U53 (x13::nil)) => P_id_U53 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) => P_id_a__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U31 (x14::x13::nil)) => P_id_U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil)) => P_id_isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_mark (measure x13) | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14:: x13::nil)) => P_id_a__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U51 (x15::x14::x13::nil)) => P_id_U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_a__isList (measure x13) | (algebra.Alg.Term algebra.F.id_isList (x13::nil)) => P_id_isList (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) => P_id_a__and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_a__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_U12 (x13::nil)) => P_id_U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_a__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_U62 (x13::nil)) => P_id_U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_a__isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isQid (x13::nil)) => P_id_isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isPal (x13::nil)) => P_id_isPal (measure x13) | (algebra.Alg.Term algebra.F.id___ (x14::x13::nil)) => P_id___ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_a__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_U43 (x13::nil)) => P_id_U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) => P_id_a__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U22 (x14::x13::nil)) => P_id_U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_a__isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_isNePal (x13::nil)) => P_id_isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U11 (x14::x13::nil)) => P_id_U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) => P_id_a__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U61 (x14::x13::nil)) => P_id_U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_a__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_U32 (x13::nil)) => P_id_U32 (measure x13) | (algebra.Alg.Term algebra.F.id_and (x14::x13::nil)) => P_id_and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o | (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) => P_id_a__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U52 (x14::x13::nil)) => P_id_U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_a__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_U23 (x13::nil)) => P_id_U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_a__isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_a__isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_isNeList (x13::nil)) => P_id_isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) => P_id_a__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U71 (x14::x13::nil)) => P_id_U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) => P_id_a__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U41 (x15::x14::x13::nil)) => P_id_U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_a__isPal (measure x13) | _ => 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____ => P_id_a____ | algebra.F.id_a => P_id_a | algebra.F.id_a__U42 => P_id_a__U42 | algebra.F.id_U42 => P_id_U42 | algebra.F.id_a__U21 => P_id_a__U21 | algebra.F.id_U21 => P_id_U21 | algebra.F.id_a__U72 => P_id_a__U72 | algebra.F.id_U72 => P_id_U72 | algebra.F.id_a__U11 => P_id_a__U11 | algebra.F.id_u => P_id_u | algebra.F.id_a__U53 => P_id_a__U53 | algebra.F.id_U53 => P_id_U53 | algebra.F.id_a__U31 => P_id_a__U31 | algebra.F.id_U31 => P_id_U31 | algebra.F.id_isPalListKind => P_id_isPalListKind | algebra.F.id_mark => P_id_mark | algebra.F.id_i => P_id_i | algebra.F.id_a__U51 => P_id_a__U51 | algebra.F.id_U51 => P_id_U51 | algebra.F.id_a__isList => P_id_a__isList | algebra.F.id_isList => P_id_isList | algebra.F.id_a__and => P_id_a__and | algebra.F.id_a__U12 => P_id_a__U12 | algebra.F.id_U12 => P_id_U12 | algebra.F.id_a__U62 => P_id_a__U62 | algebra.F.id_U62 => P_id_U62 | algebra.F.id_a__isQid => P_id_a__isQid | algebra.F.id_isQid => P_id_isQid | algebra.F.id_isPal => P_id_isPal | algebra.F.id___ => P_id___ | algebra.F.id_e => P_id_e | algebra.F.id_a__U43 => P_id_a__U43 | algebra.F.id_U43 => P_id_U43 | algebra.F.id_a__U22 => P_id_a__U22 | algebra.F.id_U22 => P_id_U22 | algebra.F.id_a__isNePal => P_id_a__isNePal | algebra.F.id_isNePal => P_id_isNePal | algebra.F.id_tt => P_id_tt | algebra.F.id_U11 => P_id_U11 | algebra.F.id_a__U61 => P_id_a__U61 | algebra.F.id_U61 => P_id_U61 | algebra.F.id_a__U32 => P_id_a__U32 | algebra.F.id_U32 => P_id_U32 | algebra.F.id_and => P_id_and | algebra.F.id_nil => P_id_nil | algebra.F.id_o => P_id_o | algebra.F.id_a__U52 => P_id_a__U52 | algebra.F.id_U52 => P_id_U52 | algebra.F.id_a__U23 => P_id_a__U23 | algebra.F.id_U23 => P_id_U23 | algebra.F.id_a__isPalListKind => P_id_a__isPalListKind | algebra.F.id_a__isNeList => P_id_a__isNeList | algebra.F.id_isNeList => P_id_isNeList | algebra.F.id_a__U71 => P_id_a__U71 | algebra.F.id_U71 => P_id_U71 | algebra.F.id_a__U41 => P_id_a__U41 | algebra.F.id_U41 => P_id_U41 | algebra.F.id_a__isPal => P_id_a__isPal 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____ => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a => match l with | nil => _ | _::_ => _ end | algebra.F.id_a__U42 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U42 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U21 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_U21 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_a__U72 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U72 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U11 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_u => match l with | nil => _ | _::_ => _ end | algebra.F.id_a__U53 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U53 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U31 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U31 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_isPalListKind => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_mark => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_i => match l with | nil => _ | _::_ => _ end | algebra.F.id_a__U51 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_U51 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_a__isList => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_isList => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__and => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U12 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U12 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U62 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U62 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__isQid => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_isQid => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_isPal => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id___ => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_e => match l with | nil => _ | _::_ => _ end | algebra.F.id_a__U43 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U43 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U22 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U22 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__isNePal => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_isNePal => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_tt => match l with | nil => _ | _::_ => _ end | algebra.F.id_U11 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U61 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U61 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U32 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U32 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_and => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_nil => match l with | nil => _ | _::_ => _ end | algebra.F.id_o => match l with | nil => _ | _::_ => _ end | algebra.F.id_a__U52 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U52 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U23 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U23 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__isPalListKind => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__isNeList => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_isNeList => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U71 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U71 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U41 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_U41 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_a__isPal => 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_____bounded;assumption. vm_compute in |-*;intros ;apply P_id_a_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U42_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U42_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U21_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U21_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U72_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U72_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U11_bounded;assumption. vm_compute in |-*;intros ;apply P_id_u_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U53_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U53_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U31_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U31_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isPalListKind_bounded;assumption. vm_compute in |-*;intros ;apply P_id_mark_bounded;assumption. vm_compute in |-*;intros ;apply P_id_i_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U51_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U51_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__and_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U12_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U12_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U62_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U62_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isQid_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isQid_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isPal_bounded;assumption. vm_compute in |-*;intros ;apply P_id____bounded;assumption. vm_compute in |-*;intros ;apply P_id_e_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U43_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U43_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U22_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U22_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNePal_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNePal_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tt_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U11_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U61_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U61_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U32_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U32_bounded;assumption. vm_compute in |-*;intros ;apply P_id_and_bounded;assumption. vm_compute in |-*;intros ;apply P_id_nil_bounded;assumption. vm_compute in |-*;intros ;apply P_id_o_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U52_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U52_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U23_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U23_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isPalListKind_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNeList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNeList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U71_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U71_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U41_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U41_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isPal_bounded;assumption. Qed. Ltac generate_pos_hyp := match goal with | H:context [measure ?x] |- _ => let v := fresh "v" in (let H := fresh "h" in (set (H:=measure_bounded x) in *;set (v:=measure x) in *; clearbody H;clearbody v)) | |- context [measure ?x] => let v := fresh "v" in (let H := fresh "h" in (set (H:=measure_bounded x) in *;set (v:=measure x) in *; clearbody H;clearbody v)) end . Hypothesis rules_monotonic : forall l r, (algebra.EQT.axiom R_xml_0_deep_rew.R_xml_0_rules r l) -> measure r <= measure l. Lemma measure_star_monotonic : forall l r, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) r l) ->measure r <= measure l. Proof. intros . do 2 (rewrite same_measure in |-*). apply InterpGen.measure_star_monotonic with (1:=Aop) (Pols:=Pols) (rules:=R_xml_0_deep_rew.R_xml_0_rules). intros f. case f. vm_compute in |-*;intros ;apply P_id_a_____monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_a__U42_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U42_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U21_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U21_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U72_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U72_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U11_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_a__U53_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U53_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U31_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U31_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isPalListKind_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_mark_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_a__U51_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U51_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isList_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isList_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__and_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U12_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U12_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U62_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U62_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isQid_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isQid_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isPal_monotonic;assumption. vm_compute in |-*;intros ;apply P_id____monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_a__U43_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U43_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U22_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U22_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isNePal_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isNePal_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_U11_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U61_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U61_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U32_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U32_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_and_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_a__U52_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U52_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U23_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U23_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isPalListKind_monotonic; assumption. vm_compute in |-*;intros ;apply P_id_a__isNeList_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isNeList_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U71_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U71_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U41_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U41_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isPal_monotonic;assumption. intros f. case f. vm_compute in |-*;intros ;apply P_id_a_____bounded;assumption. vm_compute in |-*;intros ;apply P_id_a_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U42_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U42_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U21_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U21_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U72_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U72_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U11_bounded;assumption. vm_compute in |-*;intros ;apply P_id_u_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U53_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U53_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U31_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U31_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isPalListKind_bounded;assumption. vm_compute in |-*;intros ;apply P_id_mark_bounded;assumption. vm_compute in |-*;intros ;apply P_id_i_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U51_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U51_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__and_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U12_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U12_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U62_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U62_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isQid_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isQid_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isPal_bounded;assumption. vm_compute in |-*;intros ;apply P_id____bounded;assumption. vm_compute in |-*;intros ;apply P_id_e_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U43_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U43_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U22_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U22_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNePal_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNePal_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tt_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U11_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U61_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U61_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U32_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U32_bounded;assumption. vm_compute in |-*;intros ;apply P_id_and_bounded;assumption. vm_compute in |-*;intros ;apply P_id_nil_bounded;assumption. vm_compute in |-*;intros ;apply P_id_o_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U52_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U52_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U23_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U23_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isPalListKind_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNeList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNeList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U71_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U71_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U41_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U41_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isPal_bounded;assumption. intros . do 2 (rewrite <- same_measure in |-*). apply rules_monotonic;assumption. assumption. Qed. Hypothesis P_id_A__U22 : A ->A ->A. Hypothesis P_id_A__ISNEPAL : A ->A. Hypothesis P_id_A__U43 : A ->A. Hypothesis P_id_A__U32 : A ->A. Hypothesis P_id_A__U61 : A ->A ->A. Hypothesis P_id_A__U11 : A ->A ->A. Hypothesis P_id_A__U23 : A ->A. Hypothesis P_id_A__ISPALLISTKIND : A ->A. Hypothesis P_id_A__U52 : A ->A ->A. Hypothesis P_id_A____ : A ->A ->A. Hypothesis P_id_A__U41 : A ->A ->A ->A. Hypothesis P_id_A__U71 : A ->A ->A. Hypothesis P_id_A__ISNELIST : A ->A. Hypothesis P_id_A__ISLIST : A ->A. Hypothesis P_id_A__AND : A ->A ->A. Hypothesis P_id_A__U51 : A ->A ->A ->A. Hypothesis P_id_A__ISQID : A ->A. Hypothesis P_id_A__U62 : A ->A. Hypothesis P_id_A__U12 : A ->A. Hypothesis P_id_A__U31 : A ->A ->A. Hypothesis P_id_A__ISPAL : A ->A. Hypothesis P_id_A__U53 : A ->A. Hypothesis P_id_MARK : A ->A. Hypothesis P_id_A__U42 : A ->A ->A. Hypothesis P_id_A__U72 : A ->A. Hypothesis P_id_A__U21 : A ->A ->A ->A. Hypothesis P_id_A__U22_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_A__U22 x14 x16 <= P_id_A__U22 x13 x15. Hypothesis P_id_A__ISNEPAL_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_A__ISNEPAL x14 <= P_id_A__ISNEPAL x13. Hypothesis P_id_A__U43_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_A__U43 x14 <= P_id_A__U43 x13. Hypothesis P_id_A__U32_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_A__U32 x14 <= P_id_A__U32 x13. Hypothesis P_id_A__U61_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_A__U61 x14 x16 <= P_id_A__U61 x13 x15. Hypothesis P_id_A__U11_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_A__U11 x14 x16 <= P_id_A__U11 x13 x15. Hypothesis P_id_A__U23_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_A__U23 x14 <= P_id_A__U23 x13. Hypothesis P_id_A__ISPALLISTKIND_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPALLISTKIND x14 <= P_id_A__ISPALLISTKIND x13. Hypothesis P_id_A__U52_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_A__U52 x14 x16 <= P_id_A__U52 x13 x15. Hypothesis P_id_A_____monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_A____ x14 x16 <= P_id_A____ x13 x15. Hypothesis P_id_A__U41_monotonic : forall x16 x18 x14 x17 x13 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) -> P_id_A__U41 x14 x16 x18 <= P_id_A__U41 x13 x15 x17. Hypothesis P_id_A__U71_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_A__U71 x14 x16 <= P_id_A__U71 x13 x15. Hypothesis P_id_A__ISNELIST_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNELIST x14 <= P_id_A__ISNELIST x13. Hypothesis P_id_A__ISLIST_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_A__ISLIST x14 <= P_id_A__ISLIST x13. Hypothesis P_id_A__AND_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_A__AND x14 x16 <= P_id_A__AND x13 x15. Hypothesis P_id_A__U51_monotonic : forall x16 x18 x14 x17 x13 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) -> P_id_A__U51 x14 x16 x18 <= P_id_A__U51 x13 x15 x17. Hypothesis P_id_A__ISQID_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_A__ISQID x14 <= P_id_A__ISQID x13. Hypothesis P_id_A__U62_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_A__U62 x14 <= P_id_A__U62 x13. Hypothesis P_id_A__U12_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_A__U12 x14 <= P_id_A__U12 x13. Hypothesis P_id_A__U31_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_A__U31 x14 x16 <= P_id_A__U31 x13 x15. Hypothesis P_id_A__ISPAL_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_A__ISPAL x14 <= P_id_A__ISPAL x13. Hypothesis P_id_A__U53_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_A__U53 x14 <= P_id_A__U53 x13. Hypothesis P_id_MARK_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_MARK x14 <= P_id_MARK x13. Hypothesis P_id_A__U42_monotonic : forall x16 x14 x13 x15, (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) ->P_id_A__U42 x14 x16 <= P_id_A__U42 x13 x15. Hypothesis P_id_A__U72_monotonic : forall x14 x13, (A0 <= x14)/\ (x14 <= x13) ->P_id_A__U72 x14 <= P_id_A__U72 x13. Hypothesis P_id_A__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> (A0 <= x14)/\ (x14 <= x13) -> P_id_A__U21 x14 x16 x18 <= P_id_A__U21 x13 x15 x17. Definition marked_measure t := match t with | (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) => P_id_A__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_A__ISNEPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_A__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_A__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) => P_id_A__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) => P_id_A__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_A__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_A__ISPALLISTKIND (measure x13) | (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) => P_id_A__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) => P_id_A____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) => P_id_A__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) => P_id_A__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_A__ISNELIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_A__ISLIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) => P_id_A__AND (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) => P_id_A__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_A__ISQID (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_A__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_A__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) => P_id_A__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_A__ISPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_A__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_MARK (measure x13) | (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) => P_id_A__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_A__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) => P_id_A__U21 (measure x15) (measure x14) (measure x13) | _ => 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 x13;apply (P_id_A__ISPAL x13). 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 x15 x14 x13;apply (P_id_A__U41 x15 x14 x13). 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 x14 x13;apply (P_id_A__U71 x14 x13). 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;apply (P_id_A__ISNELIST x13). 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;apply (P_id_A__ISPALLISTKIND x13). 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;apply (P_id_A__U23 x13). 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 x14 x13;apply (P_id_A__U52 x14 x13). 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;apply (P_id_A__U32 x13). 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 x14 x13;apply (P_id_A__U61 x14 x13). 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;apply (P_id_A__ISNEPAL x13). 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 x14 x13;apply (P_id_A__U22 x14 x13). 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;apply (P_id_A__U43 x13). 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;apply (P_id_A__ISQID x13). 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;apply (P_id_A__U62 x13). 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;apply (P_id_A__U12 x13). 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 x14 x13;apply (P_id_A__AND x14 x13). 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;apply (P_id_A__ISLIST x13). 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 x15 x14 x13;apply (P_id_A__U51 x15 x14 x13). 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;apply (P_id_MARK x13). 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 x14 x13;apply (P_id_A__U31 x14 x13). 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;apply (P_id_A__U53 x13). 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 x14 x13;apply (P_id_A__U11 x14 x13). 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;apply (P_id_A__U72 x13). 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 x15 x14 x13;apply (P_id_A__U21 x15 x14 x13). 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 x14 x13;apply (P_id_A__U42 x14 x13). 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 x14 x13;apply (P_id_A____ x14 x13). 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____ => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a => match l with | nil => _ | _::_ => _ end | algebra.F.id_a__U42 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U42 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U21 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_U21 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_a__U72 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U72 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U11 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_u => match l with | nil => _ | _::_ => _ end | algebra.F.id_a__U53 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U53 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U31 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U31 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_isPalListKind => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_mark => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_i => match l with | nil => _ | _::_ => _ end | algebra.F.id_a__U51 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_U51 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_a__isList => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_isList => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__and => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U12 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U12 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U62 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U62 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__isQid => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_isQid => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_isPal => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id___ => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_e => match l with | nil => _ | _::_ => _ end | algebra.F.id_a__U43 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U43 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U22 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U22 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__isNePal => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_isNePal => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_tt => match l with | nil => _ | _::_ => _ end | algebra.F.id_U11 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U61 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U61 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U32 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U32 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_and => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_nil => match l with | nil => _ | _::_ => _ end | algebra.F.id_o => match l with | nil => _ | _::_ => _ end | algebra.F.id_a__U52 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U52 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U23 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_U23 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__isPalListKind => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__isNeList => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_isNeList => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a__U71 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_U71 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_a__U41 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_U41 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_a__isPal => match l with | nil => _ | _::nil => _ | _::_::_ => _ end end. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . 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 . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . 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 . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . 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 . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . 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 . 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 . 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 . 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 . 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__U22 (x14:: x13::nil)) => P_id_A__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_A__ISNEPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_A__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_A__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_A__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_A__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_A__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_A__ISPALLISTKIND (measure x13) | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_A__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_A____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) => P_id_A__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_A__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_A__ISNELIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_A__ISLIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_A__AND (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14:: x13::nil)) => P_id_A__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_A__ISQID (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_A__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_A__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_A__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_A__ISPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_A__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_MARK (measure x13) | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_A__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_A__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14:: x13::nil)) => P_id_A__U21 (measure x15) (measure x14) (measure x13) | _ => measure t end. Proof. reflexivity . Qed. Lemma marked_measure_star_monotonic : forall f l1 l2, (closure.refl_trans_clos (closure.one_step_list (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) ) l1 l2) -> marked_measure (algebra.Alg.Term f l1) <= marked_measure (algebra.Alg.Term f l2). Proof. intros . do 2 (rewrite same_marked_measure in |-*). apply InterpGen.marked_measure_star_monotonic with (1:=Aop) (Pols:= Pols) (rules:=R_xml_0_deep_rew.R_xml_0_rules). clear f. intros f. case f. vm_compute in |-*;intros ;apply P_id_a_____monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_a__U42_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U42_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U21_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U21_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U72_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U72_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U11_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_a__U53_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U53_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U31_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U31_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isPalListKind_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_mark_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_a__U51_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U51_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isList_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isList_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__and_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U12_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U12_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U62_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U62_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isQid_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isQid_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isPal_monotonic;assumption. vm_compute in |-*;intros ;apply P_id____monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_a__U43_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U43_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U22_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U22_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isNePal_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isNePal_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_U11_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U61_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U61_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U32_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U32_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_and_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_a__U52_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U52_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U23_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U23_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isPalListKind_monotonic; assumption. vm_compute in |-*;intros ;apply P_id_a__isNeList_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_isNeList_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U71_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U71_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__U41_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_U41_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_a__isPal_monotonic;assumption. clear f. intros f. case f. vm_compute in |-*;intros ;apply P_id_a_____bounded;assumption. vm_compute in |-*;intros ;apply P_id_a_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U42_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U42_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U21_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U21_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U72_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U72_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U11_bounded;assumption. vm_compute in |-*;intros ;apply P_id_u_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U53_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U53_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U31_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U31_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isPalListKind_bounded;assumption. vm_compute in |-*;intros ;apply P_id_mark_bounded;assumption. vm_compute in |-*;intros ;apply P_id_i_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U51_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U51_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__and_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U12_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U12_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U62_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U62_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isQid_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isQid_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isPal_bounded;assumption. vm_compute in |-*;intros ;apply P_id____bounded;assumption. vm_compute in |-*;intros ;apply P_id_e_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U43_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U43_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U22_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U22_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNePal_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNePal_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tt_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U11_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U61_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U61_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U32_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U32_bounded;assumption. vm_compute in |-*;intros ;apply P_id_and_bounded;assumption. vm_compute in |-*;intros ;apply P_id_nil_bounded;assumption. vm_compute in |-*;intros ;apply P_id_o_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U52_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U52_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U23_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U23_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isPalListKind_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isNeList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_isNeList_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U71_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U71_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__U41_bounded;assumption. vm_compute in |-*;intros ;apply P_id_U41_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a__isPal_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__ISPAL_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_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__U71_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__ISNELIST_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_A__ISPALLISTKIND_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__U23_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__ISNEPAL_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_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__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_A__ISQID_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_A__U62_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__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__ISLIST_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_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_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__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__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__U72_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__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_____monotonic;assumption. discriminate H. assumption. Qed. End S. End Interp. Module InterpZ. Section S. Open Scope Z_scope. Hypothesis min_value : Z. Import ring_extention. Notation Local "'Alt'" := (Zwf.Zwf min_value). Notation Local "'Ale'" := Zle. Notation Local "'Aeq'" := (@eq Z). Notation Local "a <= b" := (Ale a b). Notation Local "a < b" := (Alt a b). Hypothesis P_id_a____ : Z ->Z ->Z. Hypothesis P_id_a : Z. Hypothesis P_id_a__U42 : Z ->Z ->Z. Hypothesis P_id_U42 : Z ->Z ->Z. Hypothesis P_id_a__U21 : Z ->Z ->Z ->Z. Hypothesis P_id_U21 : Z ->Z ->Z ->Z. Hypothesis P_id_a__U72 : Z ->Z. Hypothesis P_id_U72 : Z ->Z. Hypothesis P_id_a__U11 : Z ->Z ->Z. Hypothesis P_id_u : Z. Hypothesis P_id_a__U53 : Z ->Z. Hypothesis P_id_U53 : Z ->Z. Hypothesis P_id_a__U31 : Z ->Z ->Z. Hypothesis P_id_U31 : Z ->Z ->Z. Hypothesis P_id_isPalListKind : Z ->Z. Hypothesis P_id_mark : Z ->Z. Hypothesis P_id_i : Z. Hypothesis P_id_a__U51 : Z ->Z ->Z ->Z. Hypothesis P_id_U51 : Z ->Z ->Z ->Z. Hypothesis P_id_a__isList : Z ->Z. Hypothesis P_id_isList : Z ->Z. Hypothesis P_id_a__and : Z ->Z ->Z. Hypothesis P_id_a__U12 : Z ->Z. Hypothesis P_id_U12 : Z ->Z. Hypothesis P_id_a__U62 : Z ->Z. Hypothesis P_id_U62 : Z ->Z. Hypothesis P_id_a__isQid : Z ->Z. Hypothesis P_id_isQid : Z ->Z. Hypothesis P_id_isPal : Z ->Z. Hypothesis P_id___ : Z ->Z ->Z. Hypothesis P_id_e : Z. Hypothesis P_id_a__U43 : Z ->Z. Hypothesis P_id_U43 : Z ->Z. Hypothesis P_id_a__U22 : Z ->Z ->Z. Hypothesis P_id_U22 : Z ->Z ->Z. Hypothesis P_id_a__isNePal : Z ->Z. Hypothesis P_id_isNePal : Z ->Z. Hypothesis P_id_tt : Z. Hypothesis P_id_U11 : Z ->Z ->Z. Hypothesis P_id_a__U61 : Z ->Z ->Z. Hypothesis P_id_U61 : Z ->Z ->Z. Hypothesis P_id_a__U32 : Z ->Z. Hypothesis P_id_U32 : Z ->Z. Hypothesis P_id_and : Z ->Z ->Z. Hypothesis P_id_nil : Z. Hypothesis P_id_o : Z. Hypothesis P_id_a__U52 : Z ->Z ->Z. Hypothesis P_id_U52 : Z ->Z ->Z. Hypothesis P_id_a__U23 : Z ->Z. Hypothesis P_id_U23 : Z ->Z. Hypothesis P_id_a__isPalListKind : Z ->Z. Hypothesis P_id_a__isNeList : Z ->Z. Hypothesis P_id_isNeList : Z ->Z. Hypothesis P_id_a__U71 : Z ->Z ->Z. Hypothesis P_id_U71 : Z ->Z ->Z. Hypothesis P_id_a__U41 : Z ->Z ->Z ->Z. Hypothesis P_id_U41 : Z ->Z ->Z ->Z. Hypothesis P_id_a__isPal : Z ->Z. Hypothesis P_id_a_____monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_a____ x14 x16 <= P_id_a____ x13 x15. Hypothesis P_id_a__U42_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_a__U42 x14 x16 <= P_id_a__U42 x13 x15. Hypothesis P_id_U42_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_U42 x14 x16 <= P_id_U42 x13 x15. Hypothesis P_id_a__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_a__U21 x14 x16 x18 <= P_id_a__U21 x13 x15 x17. Hypothesis P_id_U21_monotonic : forall x16 x18 x14 x17 x13 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_U21 x14 x16 x18 <= P_id_U21 x13 x15 x17. Hypothesis P_id_a__U72_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_a__U72 x14 <= P_id_a__U72 x13. Hypothesis P_id_U72_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_U72 x14 <= P_id_U72 x13. Hypothesis P_id_a__U11_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_a__U11 x14 x16 <= P_id_a__U11 x13 x15. Hypothesis P_id_a__U53_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_a__U53 x14 <= P_id_a__U53 x13. Hypothesis P_id_U53_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_U53 x14 <= P_id_U53 x13. Hypothesis P_id_a__U31_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_a__U31 x14 x16 <= P_id_a__U31 x13 x15. Hypothesis P_id_U31_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_U31 x14 x16 <= P_id_U31 x13 x15. Hypothesis P_id_isPalListKind_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) -> P_id_isPalListKind x14 <= P_id_isPalListKind x13. Hypothesis P_id_mark_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_mark x14 <= P_id_mark x13. Hypothesis P_id_a__U51_monotonic : forall x16 x18 x14 x17 x13 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_a__U51 x14 x16 x18 <= P_id_a__U51 x13 x15 x17. Hypothesis P_id_U51_monotonic : forall x16 x18 x14 x17 x13 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_U51 x14 x16 x18 <= P_id_U51 x13 x15 x17. Hypothesis P_id_a__isList_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) -> P_id_a__isList x14 <= P_id_a__isList x13. Hypothesis P_id_isList_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_isList x14 <= P_id_isList x13. Hypothesis P_id_a__and_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_a__and x14 x16 <= P_id_a__and x13 x15. Hypothesis P_id_a__U12_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_a__U12 x14 <= P_id_a__U12 x13. Hypothesis P_id_U12_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_U12 x14 <= P_id_U12 x13. Hypothesis P_id_a__U62_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_a__U62 x14 <= P_id_a__U62 x13. Hypothesis P_id_U62_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_U62 x14 <= P_id_U62 x13. Hypothesis P_id_a__isQid_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) -> P_id_a__isQid x14 <= P_id_a__isQid x13. Hypothesis P_id_isQid_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_isQid x14 <= P_id_isQid x13. Hypothesis P_id_isPal_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_isPal x14 <= P_id_isPal x13. Hypothesis P_id____monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) ->P_id___ x14 x16 <= P_id___ x13 x15. Hypothesis P_id_a__U43_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_a__U43 x14 <= P_id_a__U43 x13. Hypothesis P_id_U43_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_U43 x14 <= P_id_U43 x13. Hypothesis P_id_a__U22_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_a__U22 x14 x16 <= P_id_a__U22 x13 x15. Hypothesis P_id_U22_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_U22 x14 x16 <= P_id_U22 x13 x15. Hypothesis P_id_a__isNePal_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) -> P_id_a__isNePal x14 <= P_id_a__isNePal x13. Hypothesis P_id_isNePal_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_isNePal x14 <= P_id_isNePal x13. Hypothesis P_id_U11_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_U11 x14 x16 <= P_id_U11 x13 x15. Hypothesis P_id_a__U61_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_a__U61 x14 x16 <= P_id_a__U61 x13 x15. Hypothesis P_id_U61_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_U61 x14 x16 <= P_id_U61 x13 x15. Hypothesis P_id_a__U32_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_a__U32 x14 <= P_id_a__U32 x13. Hypothesis P_id_U32_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_U32 x14 <= P_id_U32 x13. Hypothesis P_id_and_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_and x14 x16 <= P_id_and x13 x15. Hypothesis P_id_a__U52_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_a__U52 x14 x16 <= P_id_a__U52 x13 x15. Hypothesis P_id_U52_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_U52 x14 x16 <= P_id_U52 x13 x15. Hypothesis P_id_a__U23_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_a__U23 x14 <= P_id_a__U23 x13. Hypothesis P_id_U23_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_U23 x14 <= P_id_U23 x13. Hypothesis P_id_a__isPalListKind_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) -> P_id_a__isPalListKind x14 <= P_id_a__isPalListKind x13. Hypothesis P_id_a__isNeList_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) -> P_id_a__isNeList x14 <= P_id_a__isNeList x13. Hypothesis P_id_isNeList_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) -> P_id_isNeList x14 <= P_id_isNeList x13. Hypothesis P_id_a__U71_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_a__U71 x14 x16 <= P_id_a__U71 x13 x15. Hypothesis P_id_U71_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_U71 x14 x16 <= P_id_U71 x13 x15. Hypothesis P_id_a__U41_monotonic : forall x16 x18 x14 x17 x13 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_a__U41 x14 x16 x18 <= P_id_a__U41 x13 x15 x17. Hypothesis P_id_U41_monotonic : forall x16 x18 x14 x17 x13 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_U41 x14 x16 x18 <= P_id_U41 x13 x15 x17. Hypothesis P_id_a__isPal_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) -> P_id_a__isPal x14 <= P_id_a__isPal x13. Hypothesis P_id_a_____bounded : forall x14 x13, (min_value <= x13) -> (min_value <= x14) ->min_value <= P_id_a____ x14 x13. Hypothesis P_id_a_bounded : min_value <= P_id_a . Hypothesis P_id_a__U42_bounded : forall x14 x13, (min_value <= x13) -> (min_value <= x14) ->min_value <= P_id_a__U42 x14 x13. Hypothesis P_id_U42_bounded : forall x14 x13, (min_value <= x13) ->(min_value <= x14) ->min_value <= P_id_U42 x14 x13. Hypothesis P_id_a__U21_bounded : forall x14 x13 x15, (min_value <= x13) -> (min_value <= x14) -> (min_value <= x15) ->min_value <= P_id_a__U21 x15 x14 x13. Hypothesis P_id_U21_bounded : forall x14 x13 x15, (min_value <= x13) -> (min_value <= x14) -> (min_value <= x15) ->min_value <= P_id_U21 x15 x14 x13. Hypothesis P_id_a__U72_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_a__U72 x13. Hypothesis P_id_U72_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_U72 x13. Hypothesis P_id_a__U11_bounded : forall x14 x13, (min_value <= x13) -> (min_value <= x14) ->min_value <= P_id_a__U11 x14 x13. Hypothesis P_id_u_bounded : min_value <= P_id_u . Hypothesis P_id_a__U53_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_a__U53 x13. Hypothesis P_id_U53_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_U53 x13. Hypothesis P_id_a__U31_bounded : forall x14 x13, (min_value <= x13) -> (min_value <= x14) ->min_value <= P_id_a__U31 x14 x13. Hypothesis P_id_U31_bounded : forall x14 x13, (min_value <= x13) ->(min_value <= x14) ->min_value <= P_id_U31 x14 x13. Hypothesis P_id_isPalListKind_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_isPalListKind x13. Hypothesis P_id_mark_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_mark x13. Hypothesis P_id_i_bounded : min_value <= P_id_i . Hypothesis P_id_a__U51_bounded : forall x14 x13 x15, (min_value <= x13) -> (min_value <= x14) -> (min_value <= x15) ->min_value <= P_id_a__U51 x15 x14 x13. Hypothesis P_id_U51_bounded : forall x14 x13 x15, (min_value <= x13) -> (min_value <= x14) -> (min_value <= x15) ->min_value <= P_id_U51 x15 x14 x13. Hypothesis P_id_a__isList_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_a__isList x13. Hypothesis P_id_isList_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_isList x13. Hypothesis P_id_a__and_bounded : forall x14 x13, (min_value <= x13) -> (min_value <= x14) ->min_value <= P_id_a__and x14 x13. Hypothesis P_id_a__U12_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_a__U12 x13. Hypothesis P_id_U12_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_U12 x13. Hypothesis P_id_a__U62_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_a__U62 x13. Hypothesis P_id_U62_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_U62 x13. Hypothesis P_id_a__isQid_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_a__isQid x13. Hypothesis P_id_isQid_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_isQid x13. Hypothesis P_id_isPal_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_isPal x13. Hypothesis P_id____bounded : forall x14 x13, (min_value <= x13) ->(min_value <= x14) ->min_value <= P_id___ x14 x13. Hypothesis P_id_e_bounded : min_value <= P_id_e . Hypothesis P_id_a__U43_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_a__U43 x13. Hypothesis P_id_U43_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_U43 x13. Hypothesis P_id_a__U22_bounded : forall x14 x13, (min_value <= x13) -> (min_value <= x14) ->min_value <= P_id_a__U22 x14 x13. Hypothesis P_id_U22_bounded : forall x14 x13, (min_value <= x13) ->(min_value <= x14) ->min_value <= P_id_U22 x14 x13. Hypothesis P_id_a__isNePal_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_a__isNePal x13. Hypothesis P_id_isNePal_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_isNePal x13. Hypothesis P_id_tt_bounded : min_value <= P_id_tt . Hypothesis P_id_U11_bounded : forall x14 x13, (min_value <= x13) ->(min_value <= x14) ->min_value <= P_id_U11 x14 x13. Hypothesis P_id_a__U61_bounded : forall x14 x13, (min_value <= x13) -> (min_value <= x14) ->min_value <= P_id_a__U61 x14 x13. Hypothesis P_id_U61_bounded : forall x14 x13, (min_value <= x13) ->(min_value <= x14) ->min_value <= P_id_U61 x14 x13. Hypothesis P_id_a__U32_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_a__U32 x13. Hypothesis P_id_U32_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_U32 x13. Hypothesis P_id_and_bounded : forall x14 x13, (min_value <= x13) ->(min_value <= x14) ->min_value <= P_id_and x14 x13. Hypothesis P_id_nil_bounded : min_value <= P_id_nil . Hypothesis P_id_o_bounded : min_value <= P_id_o . Hypothesis P_id_a__U52_bounded : forall x14 x13, (min_value <= x13) -> (min_value <= x14) ->min_value <= P_id_a__U52 x14 x13. Hypothesis P_id_U52_bounded : forall x14 x13, (min_value <= x13) ->(min_value <= x14) ->min_value <= P_id_U52 x14 x13. Hypothesis P_id_a__U23_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_a__U23 x13. Hypothesis P_id_U23_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_U23 x13. Hypothesis P_id_a__isPalListKind_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_a__isPalListKind x13. Hypothesis P_id_a__isNeList_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_a__isNeList x13. Hypothesis P_id_isNeList_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_isNeList x13. Hypothesis P_id_a__U71_bounded : forall x14 x13, (min_value <= x13) -> (min_value <= x14) ->min_value <= P_id_a__U71 x14 x13. Hypothesis P_id_U71_bounded : forall x14 x13, (min_value <= x13) ->(min_value <= x14) ->min_value <= P_id_U71 x14 x13. Hypothesis P_id_a__U41_bounded : forall x14 x13 x15, (min_value <= x13) -> (min_value <= x14) -> (min_value <= x15) ->min_value <= P_id_a__U41 x15 x14 x13. Hypothesis P_id_U41_bounded : forall x14 x13 x15, (min_value <= x13) -> (min_value <= x14) -> (min_value <= x15) ->min_value <= P_id_U41 x15 x14 x13. Hypothesis P_id_a__isPal_bounded : forall x13, (min_value <= x13) ->min_value <= P_id_a__isPal x13. Definition measure := Interp.measure min_value P_id_a____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) => P_id_a____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) => P_id_a__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U42 (x14::x13::nil)) => P_id_U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14:: x13::nil)) => P_id_a__U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U21 (x15::x14::x13::nil)) => P_id_U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_a__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_U72 (x13::nil)) => P_id_U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) => P_id_a__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_a__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_U53 (x13::nil)) => P_id_U53 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) => P_id_a__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U31 (x14::x13::nil)) => P_id_U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil)) => P_id_isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_mark (measure x13) | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14:: x13::nil)) => P_id_a__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U51 (x15::x14::x13::nil)) => P_id_U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_a__isList (measure x13) | (algebra.Alg.Term algebra.F.id_isList (x13::nil)) => P_id_isList (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) => P_id_a__and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_a__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_U12 (x13::nil)) => P_id_U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_a__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_U62 (x13::nil)) => P_id_U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_a__isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isQid (x13::nil)) => P_id_isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isPal (x13::nil)) => P_id_isPal (measure x13) | (algebra.Alg.Term algebra.F.id___ (x14::x13::nil)) => P_id___ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_a__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_U43 (x13::nil)) => P_id_U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) => P_id_a__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U22 (x14::x13::nil)) => P_id_U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_a__isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_isNePal (x13::nil)) => P_id_isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U11 (x14::x13::nil)) => P_id_U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) => P_id_a__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U61 (x14::x13::nil)) => P_id_U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_a__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_U32 (x13::nil)) => P_id_U32 (measure x13) | (algebra.Alg.Term algebra.F.id_and (x14::x13::nil)) => P_id_and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o | (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) => P_id_a__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U52 (x14::x13::nil)) => P_id_U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_a__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_U23 (x13::nil)) => P_id_U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_a__isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_a__isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_isNeList (x13::nil)) => P_id_isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) => P_id_a__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U71 (x14::x13::nil)) => P_id_U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) => P_id_a__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U41 (x15::x14::x13::nil)) => P_id_U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_a__isPal (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. Qed. Hypothesis P_id_A__U22 : Z ->Z ->Z. Hypothesis P_id_A__ISNEPAL : Z ->Z. Hypothesis P_id_A__U43 : Z ->Z. Hypothesis P_id_A__U32 : Z ->Z. Hypothesis P_id_A__U61 : Z ->Z ->Z. Hypothesis P_id_A__U11 : Z ->Z ->Z. Hypothesis P_id_A__U23 : Z ->Z. Hypothesis P_id_A__ISPALLISTKIND : Z ->Z. Hypothesis P_id_A__U52 : Z ->Z ->Z. Hypothesis P_id_A____ : Z ->Z ->Z. Hypothesis P_id_A__U41 : Z ->Z ->Z ->Z. Hypothesis P_id_A__U71 : Z ->Z ->Z. Hypothesis P_id_A__ISNELIST : Z ->Z. Hypothesis P_id_A__ISLIST : Z ->Z. Hypothesis P_id_A__AND : Z ->Z ->Z. Hypothesis P_id_A__U51 : Z ->Z ->Z ->Z. Hypothesis P_id_A__ISQID : Z ->Z. Hypothesis P_id_A__U62 : Z ->Z. Hypothesis P_id_A__U12 : Z ->Z. Hypothesis P_id_A__U31 : Z ->Z ->Z. Hypothesis P_id_A__ISPAL : Z ->Z. Hypothesis P_id_A__U53 : Z ->Z. Hypothesis P_id_MARK : Z ->Z. Hypothesis P_id_A__U42 : Z ->Z ->Z. Hypothesis P_id_A__U72 : Z ->Z. Hypothesis P_id_A__U21 : Z ->Z ->Z ->Z. Hypothesis P_id_A__U22_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_A__U22 x14 x16 <= P_id_A__U22 x13 x15. Hypothesis P_id_A__ISNEPAL_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) -> P_id_A__ISNEPAL x14 <= P_id_A__ISNEPAL x13. Hypothesis P_id_A__U43_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_A__U43 x14 <= P_id_A__U43 x13. Hypothesis P_id_A__U32_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_A__U32 x14 <= P_id_A__U32 x13. Hypothesis P_id_A__U61_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_A__U61 x14 x16 <= P_id_A__U61 x13 x15. Hypothesis P_id_A__U11_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_A__U11 x14 x16 <= P_id_A__U11 x13 x15. Hypothesis P_id_A__U23_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_A__U23 x14 <= P_id_A__U23 x13. Hypothesis P_id_A__ISPALLISTKIND_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) -> P_id_A__ISPALLISTKIND x14 <= P_id_A__ISPALLISTKIND x13. Hypothesis P_id_A__U52_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_A__U52 x14 x16 <= P_id_A__U52 x13 x15. Hypothesis P_id_A_____monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_A____ x14 x16 <= P_id_A____ x13 x15. Hypothesis P_id_A__U41_monotonic : forall x16 x18 x14 x17 x13 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_A__U41 x14 x16 x18 <= P_id_A__U41 x13 x15 x17. Hypothesis P_id_A__U71_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_A__U71 x14 x16 <= P_id_A__U71 x13 x15. Hypothesis P_id_A__ISNELIST_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) -> P_id_A__ISNELIST x14 <= P_id_A__ISNELIST x13. Hypothesis P_id_A__ISLIST_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) -> P_id_A__ISLIST x14 <= P_id_A__ISLIST x13. Hypothesis P_id_A__AND_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_A__AND x14 x16 <= P_id_A__AND x13 x15. Hypothesis P_id_A__U51_monotonic : forall x16 x18 x14 x17 x13 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_A__U51 x14 x16 x18 <= P_id_A__U51 x13 x15 x17. Hypothesis P_id_A__ISQID_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) -> P_id_A__ISQID x14 <= P_id_A__ISQID x13. Hypothesis P_id_A__U62_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_A__U62 x14 <= P_id_A__U62 x13. Hypothesis P_id_A__U12_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_A__U12 x14 <= P_id_A__U12 x13. Hypothesis P_id_A__U31_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_A__U31 x14 x16 <= P_id_A__U31 x13 x15. Hypothesis P_id_A__ISPAL_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) -> P_id_A__ISPAL x14 <= P_id_A__ISPAL x13. Hypothesis P_id_A__U53_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_A__U53 x14 <= P_id_A__U53 x13. Hypothesis P_id_MARK_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_MARK x14 <= P_id_MARK x13. Hypothesis P_id_A__U42_monotonic : forall x16 x14 x13 x15, (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_A__U42 x14 x16 <= P_id_A__U42 x13 x15. Hypothesis P_id_A__U72_monotonic : forall x14 x13, (min_value <= x14)/\ (x14 <= x13) ->P_id_A__U72 x14 <= P_id_A__U72 x13. Hypothesis P_id_A__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> (min_value <= x14)/\ (x14 <= x13) -> P_id_A__U21 x14 x16 x18 <= P_id_A__U21 x13 x15 x17. Definition marked_measure := Interp.marked_measure min_value P_id_a____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal P_id_A__U22 P_id_A__ISNEPAL P_id_A__U43 P_id_A__U32 P_id_A__U61 P_id_A__U11 P_id_A__U23 P_id_A__ISPALLISTKIND P_id_A__U52 P_id_A____ P_id_A__U41 P_id_A__U71 P_id_A__ISNELIST P_id_A__ISLIST P_id_A__AND P_id_A__U51 P_id_A__ISQID P_id_A__U62 P_id_A__U12 P_id_A__U31 P_id_A__ISPAL P_id_A__U53 P_id_MARK P_id_A__U42 P_id_A__U72 P_id_A__U21 . Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_A__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_A__ISNEPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_A__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_A__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_A__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_A__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_A__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_A__ISPALLISTKIND (measure x13) | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_A__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_A____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) => P_id_A__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_A__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_A__ISNELIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_A__ISLIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_A__AND (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14:: x13::nil)) => P_id_A__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_A__ISQID (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_A__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_A__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_A__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_A__ISPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_A__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_MARK (measure x13) | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_A__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_A__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14:: x13::nil)) => P_id_A__U21 (measure x15) (measure x14) (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__ISNEPAL_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U23_monotonic;assumption. intros ;apply P_id_A__ISPALLISTKIND_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A_____monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__U71_monotonic;assumption. intros ;apply P_id_A__ISNELIST_monotonic;assumption. intros ;apply P_id_A__ISLIST_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__ISQID_monotonic;assumption. intros ;apply P_id_A__U62_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__ISPAL_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__U72_monotonic;assumption. intros ;apply P_id_A__U21_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 x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark (x1::nil))::(algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark (x2::nil))::(algebra.Alg.Term algebra.F.id_mark (x3::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_1 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_2 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark (x2::nil))::(algebra.Alg.Term algebra.F.id_mark (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_3 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x2::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_4 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x3::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_5 : forall x14 x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_nil nil) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_6 : forall x14 x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_nil nil) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_7 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_a__isNeList (x4::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) (* *) | DP_R_xml_0_8 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNeList (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) (* *) | DP_R_xml_0_9 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_10 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_11 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U23 ((algebra.Alg.Term algebra.F.id_a__isList (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) (* *) | DP_R_xml_0_12 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) (* *) | DP_R_xml_0_13 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_a__isQid (x4::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) (* *) | DP_R_xml_0_14 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isQid (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) (* *) | DP_R_xml_0_15 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_16 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_17 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_a__isNeList (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) (* *) | DP_R_xml_0_18 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNeList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) (* *) | DP_R_xml_0_19 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_20 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_21 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_a__isList (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) (* *) | DP_R_xml_0_22 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) (* *) | DP_R_xml_0_23 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U62 ((algebra.Alg.Term algebra.F.id_a__isQid (x4::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) (* *) | DP_R_xml_0_24 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isQid (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) (* *) | DP_R_xml_0_25 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U72 ((algebra.Alg.Term algebra.F.id_a__isNePal (x4::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) (* *) | DP_R_xml_0_26 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNePal (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) (* *) | DP_R_xml_0_27 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) (* *) | DP_R_xml_0_28 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_29 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_30 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5::x6::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_31 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_32 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_33 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_34 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_35 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> 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__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5::x6::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_36 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_37 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_38 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> 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__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5::x6::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_39 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_40 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_41 : forall x8 x13 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x7::(algebra.Alg.Term algebra.F.id___ (x8::x7::nil))::nil)) x13) -> 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__isQid (x7::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x7::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isPal (x8::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x8::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_42 : forall x8 x13 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x7::(algebra.Alg.Term algebra.F.id___ (x8::x7::nil))::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isQid (x7::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x7::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_43 : forall x8 x13 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x7::(algebra.Alg.Term algebra.F.id___ (x8::x7::nil))::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isQid (x7::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_44 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) (* *) | DP_R_xml_0_45 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) (* *) | DP_R_xml_0_46 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_47 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_48 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::(algebra.Alg.Term algebra.F.id_mark (x10::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_49 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_50 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_51 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_52 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_53 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_mark (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_54 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_55 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNeList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_56 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x9::x10::x11::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_57 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x9::x10::x11::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_58 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_59 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_60 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_61 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U23 ((algebra.Alg.Term algebra.F.id_mark (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_62 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_63 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_64 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_65 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_mark (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_66 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_67 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isQid (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isQid (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_68 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_69 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_70 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_71 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_72 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_mark (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_73 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_74 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x9::x10::x11::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_75 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x9::x10::x11::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_76 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_77 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_78 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_mark (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_79 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_80 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U61 (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_81 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U61 (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_82 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U62 (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U62 ((algebra.Alg.Term algebra.F.id_mark (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_83 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U62 (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_84 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_85 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_86 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__U72 ((algebra.Alg.Term algebra.F.id_mark (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_87 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_88 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNePal (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isNePal (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_89 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_90 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_91 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_92 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPal (x1::nil)) x13) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_a__isPal (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::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 x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) x13) -> DP_R_xml_0_non_scc_1 (algebra.Alg.Term algebra.F.id_a__U72 ((algebra.Alg.Term algebra.F.id_mark (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::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 x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U62 (x1::nil)) x13) -> DP_R_xml_0_non_scc_2 (algebra.Alg.Term algebra.F.id_a__U62 ((algebra.Alg.Term algebra.F.id_mark (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::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 x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x1::nil)) x13) -> DP_R_xml_0_non_scc_3 (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_mark (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::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 x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_non_scc_4 (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_mark (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::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 x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isQid (x1::nil)) x13) -> DP_R_xml_0_non_scc_5 (algebra.Alg.Term algebra.F.id_a__isQid (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::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 x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x1::nil)) x13) -> DP_R_xml_0_non_scc_6 (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_mark (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::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 x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0_non_scc_7 (algebra.Alg.Term algebra.F.id_a__U23 ((algebra.Alg.Term algebra.F.id_mark (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::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 x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x1::nil)) x13) -> DP_R_xml_0_non_scc_8 (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_mark (x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::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 x8 x13 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x7::(algebra.Alg.Term algebra.F.id___ (x8::x7::nil))::nil)) x13) -> DP_R_xml_0_non_scc_9 (algebra.Alg.Term algebra.F.id_a__isQid (x7::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::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 x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_non_scc_10 (algebra.Alg.Term algebra.F.id_a__U72 ((algebra.Alg.Term algebra.F.id_a__isNePal (x4::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::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 x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_non_scc_11 (algebra.Alg.Term algebra.F.id_a__isQid (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::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_non_scc_12 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_12_0 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_non_scc_12 (algebra.Alg.Term algebra.F.id_a__U62 ((algebra.Alg.Term algebra.F.id_a__isQid (x4::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) . Lemma acc_DP_R_xml_0_non_scc_12 : forall x y, (DP_R_xml_0_non_scc_12 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_13 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_13_0 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U61 (x9::x10::nil)) x13) -> DP_R_xml_0_non_scc_13 (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Lemma acc_DP_R_xml_0_non_scc_13 : forall x y, (DP_R_xml_0_non_scc_13 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_non_scc_12; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_11; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))). Qed. Inductive DP_R_xml_0_non_scc_14 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_14_0 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_non_scc_14 (algebra.Alg.Term algebra.F.id_a__U61 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)):: x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) . Lemma acc_DP_R_xml_0_non_scc_14 : forall x y, (DP_R_xml_0_non_scc_14 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_non_scc_12; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_11; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))). Qed. Inductive DP_R_xml_0_non_scc_15 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_15_0 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_non_scc_15 (algebra.Alg.Term algebra.F.id_a__U53 ((algebra.Alg.Term algebra.F.id_a__isList (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) . Lemma acc_DP_R_xml_0_non_scc_15 : forall x y, (DP_R_xml_0_non_scc_15 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_16 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_16_0 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_non_scc_16 (algebra.Alg.Term algebra.F.id_a__U43 ((algebra.Alg.Term algebra.F.id_a__isNeList (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) . Lemma acc_DP_R_xml_0_non_scc_16 : forall x y, (DP_R_xml_0_non_scc_16 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_17 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_17_0 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_non_scc_17 (algebra.Alg.Term algebra.F.id_a__isQid (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) . Lemma acc_DP_R_xml_0_non_scc_17 : forall x y, (DP_R_xml_0_non_scc_17 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_18 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_18_0 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_non_scc_18 (algebra.Alg.Term algebra.F.id_a__U32 ((algebra.Alg.Term algebra.F.id_a__isQid (x4::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) . Lemma acc_DP_R_xml_0_non_scc_18 : forall x y, (DP_R_xml_0_non_scc_18 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_19 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_19_0 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x9::x10::nil)) x13) -> DP_R_xml_0_non_scc_19 (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Lemma acc_DP_R_xml_0_non_scc_19 : forall x y, (DP_R_xml_0_non_scc_19 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_non_scc_18; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_17; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))). Qed. Inductive DP_R_xml_0_non_scc_20 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_20_0 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_non_scc_20 (algebra.Alg.Term algebra.F.id_a__U31 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)):: x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) . Lemma acc_DP_R_xml_0_non_scc_20 : forall x y, (DP_R_xml_0_non_scc_20 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_non_scc_18; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_17; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))). Qed. Inductive DP_R_xml_0_non_scc_21 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_21_0 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_non_scc_21 (algebra.Alg.Term algebra.F.id_a__U23 ((algebra.Alg.Term algebra.F.id_a__isList (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) . Lemma acc_DP_R_xml_0_non_scc_21 : forall x y, (DP_R_xml_0_non_scc_21 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_22 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_22_0 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_non_scc_22 (algebra.Alg.Term algebra.F.id_a__U12 ((algebra.Alg.Term algebra.F.id_a__isNeList (x4::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) . Lemma acc_DP_R_xml_0_non_scc_22 : forall x y, (DP_R_xml_0_non_scc_22 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_scc_23 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_0 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_1 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: (algebra.Alg.Term algebra.F.id_mark (x10::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_2 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark (x1::nil)):: (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark (x2::nil)):: (algebra.Alg.Term algebra.F.id_mark (x3::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_3 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark (x2::nil)):: (algebra.Alg.Term algebra.F.id_mark (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_4 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x2::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_5 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_6 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_7 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_8 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isNeList (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_9 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_10 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_11 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_12 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_13 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x1::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_14 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isNeList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_15 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5:: x6::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_16 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_17 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isNeList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_18 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_19 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_20 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_21 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5:: x6::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_22 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_23 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_24 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_25 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_26 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5:: x6::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_27 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_28 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_29 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_30 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_31 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_32 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_33 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_34 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_35 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_36 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_37 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_38 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_39 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_40 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_41 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x1::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_42 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_43 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_44 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_45 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_46 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_47 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_48 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_49 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_50 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_51 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x1::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_52 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U61 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_53 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U62 (x1::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_54 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_55 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isNePal (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_56 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_scc_23_57 : forall x8 x13 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x7::(algebra.Alg.Term algebra.F.id___ (x8::x7::nil))::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isQid (x7::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x7::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isPal (x8::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x8::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_scc_23_58 : forall x8 x13 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x7::(algebra.Alg.Term algebra.F.id___ (x8::x7::nil))::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isQid (x7::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x7::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_scc_23_59 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_60 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_61 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNePal (x1::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isNePal (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_62 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_63 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_64 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_65 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPal (x1::nil)) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isPal (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_66 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) (* *) | DP_R_xml_0_scc_23_67 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) (* *) | DP_R_xml_0_scc_23_68 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x3::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_69 : forall x14 x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_nil nil) x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_70 : forall x14 x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_nil nil) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) . Module WF_DP_R_xml_0_scc_23. Inductive DP_R_xml_0_scc_23_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_0 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_1 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::(algebra.Alg.Term algebra.F.id_mark (x10::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_2 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark (x1::nil))::(algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark (x2::nil)):: (algebra.Alg.Term algebra.F.id_mark (x3::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_3 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark (x2::nil))::(algebra.Alg.Term algebra.F.id_mark (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_4 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x2::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_5 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_6 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_7 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_8 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isNeList (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_9 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_10 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_11 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_12 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_13 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_14 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isNeList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_15 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5::x6::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_16 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_17 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isNeList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_18 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_19 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_20 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_21 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5::x6::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_22 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil)):: x6::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_23 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_24 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)):: x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_25 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_26 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5::x6::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_27 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_28 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_29 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_30 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_31 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_32 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_33 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_34 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_35 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_36 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_37 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_38 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_39 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_40 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_41 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_42 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_43 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_44 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_45 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_46 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_47 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_48 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_49 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_50 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_51 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_52 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U61 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_53 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U62 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_54 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_55 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isNePal (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_56 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_57 : forall x8 x13 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x7::(algebra.Alg.Term algebra.F.id___ (x8::x7::nil))::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isQid (x7::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x7::nil))::nil))::(algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isPal (x8::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x8::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_58 : forall x8 x13 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x7::(algebra.Alg.Term algebra.F.id___ (x8::x7::nil))::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isQid (x7::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x7::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_59 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_60 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_61 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNePal (x1::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isNePal (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_62 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_63 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_64 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_65 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPal (x1::nil)) x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isPal (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_66 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)):: x4::nil)) (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_67 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_68 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0_scc_23_large (algebra.Alg.Term algebra.F.id_mark (x3::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) . Inductive DP_R_xml_0_scc_23_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_strict_0 : forall x14 x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_nil nil) x13) -> DP_R_xml_0_scc_23_strict (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_strict_1 : forall x14 x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_nil nil) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_strict (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) . Module WF_DP_R_xml_0_scc_23_large. Inductive DP_R_xml_0_scc_23_large_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_0 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_1 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_2 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_3 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_4 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_5 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_6 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_7 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_8 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_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__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5::x6::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_9 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil)):: x6::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_10 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_11 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_12 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_13 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_14 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_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__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5::x6::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_15 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_16 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_17 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_18 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_19 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5::x6::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_20 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil)):: x6::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_21 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_22 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_23 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_24 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_25 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_26 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_27 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_28 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_29 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_30 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_31 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_32 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_33 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_34 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_35 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_36 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_37 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_38 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_39 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_40 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_41 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_42 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_43 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_44 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_45 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U61 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_46 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U62 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_47 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_48 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isNePal (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_49 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_50 : forall x8 x13 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x7::(algebra.Alg.Term algebra.F.id___ (x8::x7::nil))::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isQid (x7::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x7::nil))::nil))::(algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isPal (x8::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x8::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_51 : forall x8 x13 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x7::(algebra.Alg.Term algebra.F.id___ (x8::x7::nil))::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isQid (x7::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x7::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_52 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_53 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_54 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNePal (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isNePal (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_55 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_56 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_57 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_58 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPal (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isPal (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_59 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_60 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) . Inductive DP_R_xml_0_scc_23_large_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_strict_0 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0_scc_23_large_strict (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_strict_1 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_strict (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: (algebra.Alg.Term algebra.F.id_mark (x10::nil))::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_strict_2 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0_scc_23_large_strict (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark (x1::nil)):: (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark (x2::nil)):: (algebra.Alg.Term algebra.F.id_mark (x3::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_strict_3 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0_scc_23_large_strict (algebra.Alg.Term algebra.F.id_a____ ((algebra.Alg.Term algebra.F.id_mark (x2::nil)):: (algebra.Alg.Term algebra.F.id_mark (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_strict_4 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0_scc_23_large_strict (algebra.Alg.Term algebra.F.id_mark (x2::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_strict_5 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_strict (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_strict_6 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_strict (algebra.Alg.Term algebra.F.id_mark (x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_strict_7 : forall x2 x14 x1 x13 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x1::x2::nil)) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x13) -> DP_R_xml_0_scc_23_large_strict (algebra.Alg.Term algebra.F.id_mark (x3::nil)) (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) . Module WF_DP_R_xml_0_scc_23_large_large. Inductive DP_R_xml_0_scc_23_large_large_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_0 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_1 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_2 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_3 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_4 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_5 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_6 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_7 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_8 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_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__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5:: x6::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_9 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_10 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_11 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_12 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_13 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_14 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_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__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5:: x6::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_15 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_16 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_17 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_18 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_19 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5:: x6::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_20 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_21 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_22 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_23 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_24 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_25 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_26 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_27 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_28 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_29 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_30 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_31 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_32 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_33 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_34 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_35 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_36 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_37 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_38 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_39 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_40 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_41 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_42 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_43 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_44 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_45 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U61 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_46 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U62 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_47 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_48 : forall x8 x13 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x7::(algebra.Alg.Term algebra.F.id___ (x8::x7::nil))::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isQid (x7::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x7::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isPal (x8::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x8::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_49 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNePal (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isNePal (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_50 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_51 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_52 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Inductive DP_R_xml_0_scc_23_large_large_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_strict_0 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_strict (algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_strict_1 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_strict (algebra.Alg.Term algebra.F.id_a__isNePal (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_strict_2 : forall x8 x13 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x7::(algebra.Alg.Term algebra.F.id___ (x8::x7::nil))::nil)) x13) -> DP_R_xml_0_scc_23_large_large_strict (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isQid (x7::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x7::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_strict_3 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U71 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_strict (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_strict_4 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U72 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_strict (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_strict_5 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPal (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_strict (algebra.Alg.Term algebra.F.id_a__isPal (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_strict_6 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_strict (algebra.Alg.Term algebra.F.id_a__U71 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_strict_7 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_strict (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) . Module WF_DP_R_xml_0_scc_23_large_large_large. Inductive DP_R_xml_0_scc_23_large_large_large_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_0 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_1 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_2 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_3 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_4 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_5 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_6 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_7 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_8 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_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__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5:: x6::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_9 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_10 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_11 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_12 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_13 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_14 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_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__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5:: x6::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_15 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_16 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_17 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_18 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_19 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil))::x5:: x6::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_20 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_21 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_22 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_23 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_24 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_25 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_26 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_27 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10:: x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_28 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_29 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_30 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_31 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_32 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_33 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_34 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_35 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10:: x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_36 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_37 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_38 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_39 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_40 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10:: x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_41 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_42 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_43 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_44 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_45 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_46 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_47 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Inductive DP_R_xml_0_scc_23_large_large_large_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_strict_0 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U61 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_strict (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_strict_1 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U62 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_strict (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_strict_2 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_strict_3 : forall x8 x13 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x7::(algebra.Alg.Term algebra.F.id___ (x8::x7::nil))::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isQid (x7::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x7::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_isPal (x8::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x8::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_strict_4 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNePal (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__isNePal (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Module WF_DP_R_xml_0_scc_23_large_large_large_large. Inductive DP_R_xml_0_scc_23_large_large_large_large_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_0 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_1 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_2 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_3 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_4 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_5 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_6 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_7 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_8 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil)):: x6::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_9 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_10 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_11 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil)):: x6::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_12 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_13 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)):: x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_14 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_15 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil)):: x6::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_16 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_17 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_18 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_19 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_20 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_21 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_22 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_23 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_24 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_25 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_26 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_27 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_28 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_29 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_30 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_31 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_32 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_33 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_34 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_35 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_36 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_37 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_38 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_39 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_40 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Inductive DP_R_xml_0_scc_23_large_large_large_large_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_strict_0 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)):: x5::x6::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_strict_1 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_strict_2 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_strict_3 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__U51 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)):: x5::x6::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_strict_4 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__U21 ((algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)):: x5::x6::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_strict_5 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_strict_6 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) . Module WF_DP_R_xml_0_scc_23_large_large_large_large_large. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_0 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U11 (( algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_1 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_2 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_3 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__and (( algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: ( algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_4 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_5 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_6 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_7 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_8 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_9 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isNeList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_10 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_11 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_12 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U11 (( algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)):: x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_13 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_14 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_15 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_16 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U22 (( algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_17 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_18 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_19 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_20 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_21 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_22 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U41 (( algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10:: x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_23 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_24 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__U42 (( algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_25 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_26 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_27 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_28 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__and (( algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_29 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_30 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Inductive DP_R_xml_0_scc_23_large_large_large_large_large_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_strict_0 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_strict_1 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_strict_2 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_strict_3 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_strict_4 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__U21 (( algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10:: x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_strict_5 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U21 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_strict_6 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__U51 (( algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10:: x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_strict_7 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U51 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_strict_8 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_a__U52 (( algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_strict_9 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U52 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_strict (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Module WF_DP_R_xml_0_scc_23_large_large_large_large_large_large. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_non_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_non_scc_1_0 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_non_scc_1 (algebra.Alg.Term algebra.F.id_a__U52 ((algebra.Alg.Term algebra.F.id_a__isNeList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) . Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_non_scc_1 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_non_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large.DP_R_xml_0_scc_23_large_large_large_large_large_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_23_large_large_large_large_large_large_scc_2 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_0 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNeList (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_1 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_2 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_3 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_4 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_5 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_6 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_7 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNeList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_8 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_9 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_10 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_11 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_12 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_13 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_14 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_15 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_16 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_17 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_18 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_19 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_20 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isNeList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_21 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_22 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_23 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_24 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_25 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_26 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_27 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_28 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_29 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Module WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_0 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_1 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_2 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_3 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_4 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNeList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_5 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_6 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_7 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_8 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_9 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_10 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_11 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_12 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_13 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_14 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_15 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isNeList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_16 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_17 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_18 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_19 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_20 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_21 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_22 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_23 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_strict_0 : forall x4 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_strict (algebra.Alg.Term algebra.F.id_a__isNeList (x4::nil)) (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_strict_1 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_strict (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_strict_2 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U12 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_strict (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_strict_3 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_strict (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_strict_4 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_strict (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_strict_5 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_strict (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Module WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_1_0 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_1 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil))::x4::nil)) (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) . Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_1 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_2 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_2_0 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_2 (algebra.Alg.Term algebra.F.id_a__isList (x5::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) . Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_2 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_2 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_3 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_3_0 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_3 (algebra.Alg.Term algebra.F.id_a__isList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_3 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_3 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_4 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_4_0 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_4 (algebra.Alg.Term algebra.F.id_a__isList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) . Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_4 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_4 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_5 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_5_0 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U22 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_5 (algebra.Alg.Term algebra.F.id_a__U22 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_5 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_5 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_large_non_scc_6 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_6_0 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U11 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_6 (algebra.Alg.Term algebra.F.id_a__U11 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_6 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_6 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_0 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_1 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_2 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_a__isNeList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_3 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_4 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_5 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_6 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_7 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_8 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_9 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_10 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_a__isNeList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_11 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_12 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_13 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_14 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_15 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_16 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_17 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Module WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_0 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_1 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_2 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_3 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_4 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_5 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_6 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_7 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_8 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_9 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_10 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_11 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_strict_0 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isNeList (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_strict (algebra.Alg.Term algebra.F.id_a__isNeList (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_strict_1 : forall x4 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_strict (algebra.Alg.Term algebra.F.id_a__isPalListKind (x4::nil)) (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_strict_2 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U31 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_strict (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_strict_3 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U32 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_strict (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_strict_4 : forall x6 x14 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_strict (algebra.Alg.Term algebra.F.id_a__isNeList (x6::nil)) (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_strict_5 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_strict (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Module WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_1_0 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U42 (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_1 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_1 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_2 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_2_0 : forall x6 x14 x5 x13 x15, (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) x15) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x6 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_2 (algebra.Alg.Term algebra.F.id_a__U42 ((algebra.Alg.Term algebra.F.id_a__isList (x5::nil)):: x6::nil)) (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) . Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_2 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_2 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_3 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_3_0 : forall x10 x9 x13 x11, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U41 (x9::x10::x11::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_3 (algebra.Alg.Term algebra.F.id_a__U41 ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::x11::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_3 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_3 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_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' ))). Qed. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_0 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_1 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_2 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_3 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_4 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_5 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_6 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_7 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_8 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) . Module WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_0 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_1 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_2 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_3 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil))::x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_4 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_5 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_6 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)):: (algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_7 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) . Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_strict_0 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U53 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_strict (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Module WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_0 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_1 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_2 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_3 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_4 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_5 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_6 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) . Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_strict_0 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U23 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_strict (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Module WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_0 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_1 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_2 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_3 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_4 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_strict_0 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_strict (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil))::(algebra.Alg.Term algebra.F.id_isPalListKind (x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_strict_1 : forall x6 x5 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id___ (x5::x6::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_strict (algebra.Alg.Term algebra.F.id_a__isPalListKind (x5::nil)) (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) . Module WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_non_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_non_scc_1_0 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_isPalListKind (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_non_scc_1 (algebra.Alg.Term algebra.F.id_a__isPalListKind (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_non_scc_1 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_non_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_0 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_1 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_2 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_3 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2 (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Module WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_0 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_1 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) . Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_strict_0 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_strict (algebra.Alg.Term algebra.F.id_a__and ((algebra.Alg.Term algebra.F.id_mark (x9::nil)):: x10::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_strict_1 : forall x10 x9 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_and (x9::x10::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_strict (algebra.Alg.Term algebra.F.id_mark (x9::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Module WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_scc_1_0 : forall x1 x13, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_U43 (x1::nil)) x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_scc_1 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_mark (x13::nil)) . Module WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_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____ (x13:Z) (x14:Z) := 2 + 2* x13 + 1* x14. Definition P_id_a := 0. Definition P_id_a__U42 (x13:Z) (x14:Z) := 1 + 2* x14. Definition P_id_U42 (x13:Z) (x14:Z) := 1 + 2* x14. Definition P_id_a__U21 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_U21 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_a__U72 (x13:Z) := 0. Definition P_id_U72 (x13:Z) := 0. Definition P_id_a__U11 (x13:Z) (x14:Z) := 0. Definition P_id_u := 0. Definition P_id_a__U53 (x13:Z) := 0. Definition P_id_U53 (x13:Z) := 0. Definition P_id_a__U31 (x13:Z) (x14:Z) := 2* x14. Definition P_id_U31 (x13:Z) (x14:Z) := 2* x14. Definition P_id_isPalListKind (x13:Z) := 0. Definition P_id_mark (x13:Z) := 1* x13. Definition P_id_i := 0. Definition P_id_a__U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_a__isList (x13:Z) := 0. Definition P_id_isList (x13:Z) := 0. Definition P_id_a__and (x13:Z) (x14:Z) := 2* x14. Definition P_id_a__U12 (x13:Z) := 0. Definition P_id_U12 (x13:Z) := 0. Definition P_id_a__U62 (x13:Z) := 0. Definition P_id_U62 (x13:Z) := 0. Definition P_id_a__isQid (x13:Z) := 0. Definition P_id_isQid (x13:Z) := 0. Definition P_id_isPal (x13:Z) := 0. Definition P_id___ (x13:Z) (x14:Z) := 2 + 2* x13 + 1* x14 . Definition P_id_e := 0. Definition P_id_a__U43 (x13:Z) := 1 + 1* x13. Definition P_id_U43 (x13:Z) := 1 + 1* x13. Definition P_id_a__U22 (x13:Z) (x14:Z) := 0. Definition P_id_U22 (x13:Z) (x14:Z) := 0. Definition P_id_a__isNePal (x13:Z) := 0. Definition P_id_isNePal (x13:Z) := 0. Definition P_id_tt := 0. Definition P_id_U11 (x13:Z) (x14:Z) := 0. Definition P_id_a__U61 (x13:Z) (x14:Z) := 0. Definition P_id_U61 (x13:Z) (x14:Z) := 0. Definition P_id_a__U32 (x13:Z) := 0. Definition P_id_U32 (x13:Z) := 0. Definition P_id_and (x13:Z) (x14:Z) := 2* x14. Definition P_id_nil := 0. Definition P_id_o := 0. Definition P_id_a__U52 (x13:Z) (x14:Z) := 0. Definition P_id_U52 (x13:Z) (x14:Z) := 0. Definition P_id_a__U23 (x13:Z) := 0. Definition P_id_U23 (x13:Z) := 0. Definition P_id_a__isPalListKind (x13:Z) := 0. Definition P_id_a__isNeList (x13:Z) := 2* x13. Definition P_id_isNeList (x13:Z) := 2* x13. Definition P_id_a__U71 (x13:Z) (x14:Z) := 0. Definition P_id_U71 (x13:Z) (x14:Z) := 0. Definition P_id_a__U41 (x13:Z) (x14:Z) (x15:Z) := 1 + 2* x15. Definition P_id_U41 (x13:Z) (x14:Z) (x15:Z) := 1 + 2* x15 . Definition P_id_a__isPal (x13:Z) := 0. Lemma P_id_a_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a____ x14 x16 <= P_id_a____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U42 x14 x16 <= P_id_a__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U42 x14 x16 <= P_id_U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U21 x14 x16 x18 <= P_id_a__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U21 x14 x16 x18 <= P_id_U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U72 x14 <= P_id_a__U72 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_U72 x14 <= P_id_U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U11 x14 x16 <= P_id_a__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U53 x14 <= P_id_a__U53 x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_U53 x14 <= P_id_U53 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U31 x14 x16 <= P_id_a__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U31 x14 x16 <= P_id_U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isPalListKind x14 <= P_id_isPalListKind x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_mark x14 <= P_id_mark x13. Proof. intros x14 x13. intros [H_1 H_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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U51 x14 x16 x18 <= P_id_a__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U51_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U51 x14 x16 x18 <= P_id_U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isList x14 <= P_id_a__isList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isList x14 <= P_id_isList x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__and x14 x16 <= P_id_a__and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U12_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U12 x14 <= P_id_a__U12 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_U12 x14 <= P_id_U12 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U62 x14 <= P_id_a__U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_U62 x14 <= P_id_U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isQid x14 <= P_id_a__isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isQid x14 <= P_id_isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isPal x14 <= P_id_isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id___ x14 x16 <= P_id___ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U43_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U43 x14 <= P_id_a__U43 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_U43 x14 <= P_id_U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U22 x14 x16 <= P_id_a__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U22 x14 x16 <= P_id_U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNePal x14 <= P_id_a__isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isNePal x14 <= P_id_isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U11_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U11 x14 x16 <= P_id_U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U61 x14 x16 <= P_id_a__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U61 x14 x16 <= P_id_U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U32 x14 <= P_id_a__U32 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_U32 x14 <= P_id_U32 x13. Proof. intros x14 x13. intros [H_1 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_and x14 x16 <= P_id_and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U52 x14 x16 <= P_id_a__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U52 x14 x16 <= P_id_U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U23 x14 <= P_id_a__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_U23 x14 <= P_id_U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPalListKind x14 <= P_id_a__isPalListKind x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNeList x14 <= P_id_a__isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isNeList x14 <= P_id_isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U71 x14 x16 <= P_id_a__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U71 x14 x16 <= P_id_U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U41 x14 x16 x18 <= P_id_a__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U41_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U41 x14 x16 x18 <= P_id_U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPal x14 <= P_id_a__isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a____ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U42 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U42 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U21 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U21 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U72 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_U72 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U11 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_u_bounded : 0 <= P_id_u . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U53 x13. 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 x13, (0 <= x13) ->0 <= P_id_U53 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U31 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U31 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPalListKind x13. 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 x13, (0 <= x13) ->0 <= P_id_mark x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_i_bounded : 0 <= P_id_i . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_bounded : forall x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U51 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U51 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isList x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__and x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U12 x13. 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 x13, (0 <= x13) ->0 <= P_id_U12 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id___ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_e_bounded : 0 <= P_id_e . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U43 x13. 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 x13, (0 <= x13) ->0 <= P_id_U43 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U22 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U22 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNePal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNePal x13. 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_U11_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U11 x14 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U61 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U61 x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U32 x13. 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 x13, (0 <= x13) ->0 <= P_id_U32 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_and x14 x13. 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_o_bounded : 0 <= P_id_o . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U52 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U52 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPalListKind x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U71 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U71 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U41 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U41 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPal x13. 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____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) => P_id_a____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) => P_id_a__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U42 (x14::x13::nil)) => P_id_U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) => P_id_a__U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U21 (x15::x14::x13::nil)) => P_id_U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_a__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_U72 (x13::nil)) => P_id_U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) => P_id_a__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_a__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_U53 (x13::nil)) => P_id_U53 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) => P_id_a__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U31 (x14::x13::nil)) => P_id_U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil)) => P_id_isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_mark (measure x13) | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) => P_id_a__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U51 (x15::x14::x13::nil)) => P_id_U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_a__isList (measure x13) | (algebra.Alg.Term algebra.F.id_isList (x13::nil)) => P_id_isList (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) => P_id_a__and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_a__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_U12 (x13::nil)) => P_id_U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_a__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_U62 (x13::nil)) => P_id_U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_a__isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isQid (x13::nil)) => P_id_isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isPal (x13::nil)) => P_id_isPal (measure x13) | (algebra.Alg.Term algebra.F.id___ (x14::x13::nil)) => P_id___ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_a__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_U43 (x13::nil)) => P_id_U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) => P_id_a__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U22 (x14::x13::nil)) => P_id_U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_a__isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_isNePal (x13::nil)) => P_id_isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U11 (x14::x13::nil)) => P_id_U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) => P_id_a__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U61 (x14::x13::nil)) => P_id_U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_a__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_U32 (x13::nil)) => P_id_U32 (measure x13) | (algebra.Alg.Term algebra.F.id_and (x14::x13::nil)) => P_id_and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o | (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) => P_id_a__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U52 (x14::x13::nil)) => P_id_U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_a__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_U23 (x13::nil)) => P_id_U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_a__isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_a__isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_isNeList (x13::nil)) => P_id_isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) => P_id_a__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U71 (x14::x13::nil)) => P_id_U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) => P_id_a__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U41 (x15::x14::x13::nil)) => P_id_U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_a__isPal (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic; assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__U22 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNEPAL (x13:Z) := 0. Definition P_id_A__U43 (x13:Z) := 0. Definition P_id_A__U32 (x13:Z) := 0. Definition P_id_A__U61 (x13:Z) (x14:Z) := 0. Definition P_id_A__U11 (x13:Z) (x14:Z) := 0. Definition P_id_A__U23 (x13:Z) := 0. Definition P_id_A__ISPALLISTKIND (x13:Z) := 0. Definition P_id_A__U52 (x13:Z) (x14:Z) := 0. Definition P_id_A____ (x13:Z) (x14:Z) := 0. Definition P_id_A__U41 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_A__U71 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNELIST (x13:Z) := 0. Definition P_id_A__ISLIST (x13:Z) := 0. Definition P_id_A__AND (x13:Z) (x14:Z) := 0. Definition P_id_A__U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_A__ISQID (x13:Z) := 0. Definition P_id_A__U62 (x13:Z) := 0. Definition P_id_A__U12 (x13:Z) := 0. Definition P_id_A__U31 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISPAL (x13:Z) := 0. Definition P_id_A__U53 (x13:Z) := 0. Definition P_id_MARK (x13:Z) := 1* x13. Definition P_id_A__U42 (x13:Z) (x14:Z) := 0. Definition P_id_A__U72 (x13:Z) := 0. Definition P_id_A__U21 (x13:Z) (x14:Z) (x15:Z) := 0. Lemma P_id_A__U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U22 x14 x16 <= P_id_A__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNEPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNEPAL x14 <= P_id_A__ISNEPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U43 x14 <= P_id_A__U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U32 x14 <= P_id_A__U32 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U61 x14 x16 <= P_id_A__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U11 x14 x16 <= P_id_A__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U23 x14 <= P_id_A__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISPALLISTKIND_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPALLISTKIND x14 <= P_id_A__ISPALLISTKIND x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U52 x14 x16 <= P_id_A__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A____ x14 x16 <= P_id_A____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U41 x14 x16 x18 <= P_id_A__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U71 x14 x16 <= P_id_A__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNELIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNELIST x14 <= P_id_A__ISNELIST x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISLIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISLIST x14 <= P_id_A__ISLIST x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__AND x14 x16 <= P_id_A__AND x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U51 x14 x16 x18 <= P_id_A__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__ISQID_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISQID x14 <= P_id_A__ISQID x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U62 x14 <= P_id_A__U62 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U12 x14 <= P_id_A__U12 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U31 x14 x16 <= P_id_A__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPAL x14 <= P_id_A__ISPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U53 x14 <= P_id_A__U53 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_MARK x14 <= P_id_MARK x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U42 x14 x16 <= P_id_A__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U72 x14 <= P_id_A__U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U21 x14 x16 x18 <= P_id_A__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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. Definition marked_measure := InterpZ.marked_measure 0 P_id_a____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal P_id_A__U22 P_id_A__ISNEPAL P_id_A__U43 P_id_A__U32 P_id_A__U61 P_id_A__U11 P_id_A__U23 P_id_A__ISPALLISTKIND P_id_A__U52 P_id_A____ P_id_A__U41 P_id_A__U71 P_id_A__ISNELIST P_id_A__ISLIST P_id_A__AND P_id_A__U51 P_id_A__ISQID P_id_A__U62 P_id_A__U12 P_id_A__U31 P_id_A__ISPAL P_id_A__U53 P_id_MARK P_id_A__U42 P_id_A__U72 P_id_A__U21. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_A__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_A__ISNEPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_A__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_A__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_A__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_A__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_A__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_A__ISPALLISTKIND (measure x13) | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_A__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_A____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) => P_id_A__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_A__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_A__ISNELIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_A__ISLIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_A__AND (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14:: x13::nil)) => P_id_A__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_A__ISQID (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_A__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_A__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_A__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_A__ISPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_A__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_MARK (measure x13) | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_A__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_A__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14:: x13::nil)) => P_id_A__U21 (measure x15) (measure x14) (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic; assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__ISNEPAL_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U23_monotonic;assumption. intros ;apply P_id_A__ISPALLISTKIND_monotonic; assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A_____monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__U71_monotonic;assumption. intros ;apply P_id_A__ISNELIST_monotonic;assumption. intros ;apply P_id_A__ISLIST_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__ISQID_monotonic;assumption. intros ;apply P_id_A__U62_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__ISPAL_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__U72_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Lemma wf : well_founded WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_scc_1. Definition wf_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_scc_1 := WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_scc_1.wf . Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_scc_1 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_scc_1. Qed. Inductive DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_non_scc_2 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_non_scc_2_0 : forall x14 x1 x13, (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) x14) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x13) -> DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_non_scc_2 (algebra.Alg.Term algebra.F.id_mark (x1::nil)) (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) . Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_non_scc_2 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_non_scc_2 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large . Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_non_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (fail)))))). Qed. End WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_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____ (x13:Z) (x14:Z) := 1 + 3* x13 + 1* x14. Definition P_id_a := 0. Definition P_id_a__U42 (x13:Z) (x14:Z) := 0. Definition P_id_U42 (x13:Z) (x14:Z) := 0. Definition P_id_a__U21 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_U21 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_a__U72 (x13:Z) := 0. Definition P_id_U72 (x13:Z) := 0. Definition P_id_a__U11 (x13:Z) (x14:Z) := 0. Definition P_id_u := 3. Definition P_id_a__U53 (x13:Z) := 0. Definition P_id_U53 (x13:Z) := 0. Definition P_id_a__U31 (x13:Z) (x14:Z) := 0. Definition P_id_U31 (x13:Z) (x14:Z) := 0. Definition P_id_isPalListKind (x13:Z) := 1* x13. Definition P_id_mark (x13:Z) := 1* x13. Definition P_id_i := 0. Definition P_id_a__U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_a__isList (x13:Z) := 0. Definition P_id_isList (x13:Z) := 0. Definition P_id_a__and (x13:Z) (x14:Z) := 1 + 2* x13 + 1* x14. Definition P_id_a__U12 (x13:Z) := 0. Definition P_id_U12 (x13:Z) := 0. Definition P_id_a__U62 (x13:Z) := 0. Definition P_id_U62 (x13:Z) := 0. Definition P_id_a__isQid (x13:Z) := 0. Definition P_id_isQid (x13:Z) := 0. Definition P_id_isPal (x13:Z) := 1 + 2* x13. Definition P_id___ (x13:Z) (x14:Z) := 1 + 3* x13 + 1* x14. Definition P_id_e := 0. Definition P_id_a__U43 (x13:Z) := 2* x13. Definition P_id_U43 (x13:Z) := 2* x13. Definition P_id_a__U22 (x13:Z) (x14:Z) := 0. Definition P_id_U22 (x13:Z) (x14:Z) := 0. Definition P_id_a__isNePal (x13:Z) := 2 + 2* x13. Definition P_id_isNePal (x13:Z) := 2 + 2* x13. Definition P_id_tt := 0. Definition P_id_U11 (x13:Z) (x14:Z) := 0. Definition P_id_a__U61 (x13:Z) (x14:Z) := 1* x14. Definition P_id_U61 (x13:Z) (x14:Z) := 1* x14. Definition P_id_a__U32 (x13:Z) := 0. Definition P_id_U32 (x13:Z) := 0. Definition P_id_and (x13:Z) (x14:Z) := 1 + 2* x13 + 1* x14. Definition P_id_nil := 0. Definition P_id_o := 0. Definition P_id_a__U52 (x13:Z) (x14:Z) := 0. Definition P_id_U52 (x13:Z) (x14:Z) := 0. Definition P_id_a__U23 (x13:Z) := 0. Definition P_id_U23 (x13:Z) := 0. Definition P_id_a__isPalListKind (x13:Z) := 1* x13. Definition P_id_a__isNeList (x13:Z) := 0. Definition P_id_isNeList (x13:Z) := 0. Definition P_id_a__U71 (x13:Z) (x14:Z) := 0. Definition P_id_U71 (x13:Z) (x14:Z) := 0. Definition P_id_a__U41 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_U41 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_a__isPal (x13:Z) := 1 + 2* x13. Lemma P_id_a_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a____ x14 x16 <= P_id_a____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U42 x14 x16 <= P_id_a__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U42 x14 x16 <= P_id_U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U21 x14 x16 x18 <= P_id_a__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U21 x14 x16 x18 <= P_id_U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U72 x14 <= P_id_a__U72 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U72 x14 <= P_id_U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U11 x14 x16 <= P_id_a__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U53 x14 <= P_id_a__U53 x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U53 x14 <= P_id_U53 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U31 x14 x16 <= P_id_a__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U31 x14 x16 <= P_id_U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isPalListKind x14 <= P_id_isPalListKind x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_mark x14 <= P_id_mark x13. Proof. intros x14 x13. intros [H_1 H_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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U51 x14 x16 x18 <= P_id_a__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U51_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U51 x14 x16 x18 <= P_id_U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isList x14 <= P_id_a__isList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isList x14 <= P_id_isList x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__and x14 x16 <= P_id_a__and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U12_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U12 x14 <= P_id_a__U12 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U12 x14 <= P_id_U12 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U62 x14 <= P_id_a__U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U62 x14 <= P_id_U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isQid x14 <= P_id_a__isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isQid x14 <= P_id_isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isPal x14 <= P_id_isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id___ x14 x16 <= P_id___ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U43_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U43 x14 <= P_id_a__U43 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U43 x14 <= P_id_U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U22 x14 x16 <= P_id_a__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U22 x14 x16 <= P_id_U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNePal x14 <= P_id_a__isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isNePal x14 <= P_id_isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U11_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U11 x14 x16 <= P_id_U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U61 x14 x16 <= P_id_a__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U61 x14 x16 <= P_id_U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U32 x14 <= P_id_a__U32 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U32 x14 <= P_id_U32 x13. Proof. intros x14 x13. intros [H_1 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_and x14 x16 <= P_id_and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U52 x14 x16 <= P_id_a__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U52 x14 x16 <= P_id_U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U23 x14 <= P_id_a__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U23 x14 <= P_id_U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPalListKind x14 <= P_id_a__isPalListKind x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNeList x14 <= P_id_a__isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isNeList x14 <= P_id_isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U71 x14 x16 <= P_id_a__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U71 x14 x16 <= P_id_U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U41 x14 x16 x18 <= P_id_a__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U41_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U41 x14 x16 x18 <= P_id_U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPal x14 <= P_id_a__isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a____ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U42 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U42 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U21 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U21 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U72 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_U72 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U11 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_u_bounded : 0 <= P_id_u . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U53 x13. 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 x13, (0 <= x13) ->0 <= P_id_U53 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U31 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U31 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPalListKind x13. 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 x13, (0 <= x13) ->0 <= P_id_mark x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_i_bounded : 0 <= P_id_i . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_bounded : forall x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U51 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U51 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isList x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__and x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U12 x13. 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 x13, (0 <= x13) ->0 <= P_id_U12 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id___ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_e_bounded : 0 <= P_id_e . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U43 x13. 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 x13, (0 <= x13) ->0 <= P_id_U43 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U22 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U22 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNePal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNePal x13. 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_U11_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U11 x14 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U61 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U61 x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U32 x13. 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 x13, (0 <= x13) ->0 <= P_id_U32 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_and x14 x13. 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_o_bounded : 0 <= P_id_o . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U52 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U52 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPalListKind x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U71 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U71 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U41 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U41 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPal x13. 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____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) => P_id_a____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) => P_id_a__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U42 (x14:: x13::nil)) => P_id_U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) => P_id_a__U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U21 (x15:: x14::x13::nil)) => P_id_U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_a__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_U72 (x13::nil)) => P_id_U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) => P_id_a__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_a__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_U53 (x13::nil)) => P_id_U53 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) => P_id_a__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U31 (x14:: x13::nil)) => P_id_U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil)) => P_id_isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_mark (measure x13) | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) => P_id_a__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U51 (x15:: x14::x13::nil)) => P_id_U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_a__isList (measure x13) | (algebra.Alg.Term algebra.F.id_isList (x13::nil)) => P_id_isList (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) => P_id_a__and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_a__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_U12 (x13::nil)) => P_id_U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_a__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_U62 (x13::nil)) => P_id_U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_a__isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isQid (x13::nil)) => P_id_isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isPal (x13::nil)) => P_id_isPal (measure x13) | (algebra.Alg.Term algebra.F.id___ (x14:: x13::nil)) => P_id___ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_a__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_U43 (x13::nil)) => P_id_U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) => P_id_a__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U22 (x14:: x13::nil)) => P_id_U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_a__isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_isNePal (x13::nil)) => P_id_isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U11 (x14:: x13::nil)) => P_id_U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) => P_id_a__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U61 (x14:: x13::nil)) => P_id_U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_a__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_U32 (x13::nil)) => P_id_U32 (measure x13) | (algebra.Alg.Term algebra.F.id_and (x14:: x13::nil)) => P_id_and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o | (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) => P_id_a__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U52 (x14:: x13::nil)) => P_id_U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_a__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_U23 (x13::nil)) => P_id_U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_a__isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_a__isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_isNeList (x13::nil)) => P_id_isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) => P_id_a__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U71 (x14:: x13::nil)) => P_id_U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) => P_id_a__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U41 (x15:: x14::x13::nil)) => P_id_U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_a__isPal (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__U22 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNEPAL (x13:Z) := 0. Definition P_id_A__U43 (x13:Z) := 0. Definition P_id_A__U32 (x13:Z) := 0. Definition P_id_A__U61 (x13:Z) (x14:Z) := 0. Definition P_id_A__U11 (x13:Z) (x14:Z) := 0. Definition P_id_A__U23 (x13:Z) := 0. Definition P_id_A__ISPALLISTKIND (x13:Z) := 0. Definition P_id_A__U52 (x13:Z) (x14:Z) := 0. Definition P_id_A____ (x13:Z) (x14:Z) := 0. Definition P_id_A__U41 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_A__U71 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNELIST (x13:Z) := 0. Definition P_id_A__ISLIST (x13:Z) := 0. Definition P_id_A__AND (x13:Z) (x14:Z) := 1* x14. Definition P_id_A__U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_A__ISQID (x13:Z) := 0. Definition P_id_A__U62 (x13:Z) := 0. Definition P_id_A__U12 (x13:Z) := 0. Definition P_id_A__U31 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISPAL (x13:Z) := 0. Definition P_id_A__U53 (x13:Z) := 0. Definition P_id_MARK (x13:Z) := 1* x13. Definition P_id_A__U42 (x13:Z) (x14:Z) := 0. Definition P_id_A__U72 (x13:Z) := 0. Definition P_id_A__U21 (x13:Z) (x14:Z) (x15:Z) := 0. Lemma P_id_A__U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U22 x14 x16 <= P_id_A__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNEPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNEPAL x14 <= P_id_A__ISNEPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U43 x14 <= P_id_A__U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U32 x14 <= P_id_A__U32 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U61 x14 x16 <= P_id_A__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U11 x14 x16 <= P_id_A__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U23 x14 <= P_id_A__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISPALLISTKIND_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPALLISTKIND x14 <= P_id_A__ISPALLISTKIND x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U52 x14 x16 <= P_id_A__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A____ x14 x16 <= P_id_A____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U41 x14 x16 x18 <= P_id_A__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U71 x14 x16 <= P_id_A__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNELIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNELIST x14 <= P_id_A__ISNELIST x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISLIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISLIST x14 <= P_id_A__ISLIST x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__AND x14 x16 <= P_id_A__AND x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U51 x14 x16 x18 <= P_id_A__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__ISQID_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISQID x14 <= P_id_A__ISQID x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)|| (repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U62 x14 <= P_id_A__U62 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U12 x14 <= P_id_A__U12 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U31 x14 x16 <= P_id_A__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPAL x14 <= P_id_A__ISPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U53 x14 <= P_id_A__U53 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_MARK x14 <= P_id_MARK x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U42 x14 x16 <= P_id_A__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U72 x14 <= P_id_A__U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U21 x14 x16 x18 <= P_id_A__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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. Definition marked_measure := InterpZ.marked_measure 0 P_id_a____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal P_id_A__U22 P_id_A__ISNEPAL P_id_A__U43 P_id_A__U32 P_id_A__U61 P_id_A__U11 P_id_A__U23 P_id_A__ISPALLISTKIND P_id_A__U52 P_id_A____ P_id_A__U41 P_id_A__U71 P_id_A__ISNELIST P_id_A__ISLIST P_id_A__AND P_id_A__U51 P_id_A__ISQID P_id_A__U62 P_id_A__U12 P_id_A__U31 P_id_A__ISPAL P_id_A__U53 P_id_MARK P_id_A__U42 P_id_A__U72 P_id_A__U21. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_A__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_A__ISNEPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_A__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_A__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_A__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_A__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_A__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_A__ISPALLISTKIND (measure x13) | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_A__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_A____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) => P_id_A__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_A__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_A__ISNELIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_A__ISLIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_A__AND (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14:: x13::nil)) => P_id_A__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_A__ISQID (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_A__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_A__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_A__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_A__ISPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_A__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_MARK (measure x13) | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_A__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_A__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14:: x13::nil)) => P_id_A__U21 (measure x15) (measure x14) (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__ISNEPAL_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U23_monotonic;assumption. intros ;apply P_id_A__ISPALLISTKIND_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A_____monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__U71_monotonic;assumption. intros ;apply P_id_A__ISNELIST_monotonic;assumption. intros ;apply P_id_A__ISLIST_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__ISQID_monotonic;assumption. intros ;apply P_id_A__U62_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__ISPAL_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__U72_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b). Definition le a b := marked_measure a <= marked_measure b. Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c. Proof. unfold lt, le in *. intros a b c. apply (interp.le_lt_compat_right (interp.o_Z 0)). Qed. Lemma wf_lt : well_founded lt. Proof. unfold lt in *. apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. Qed. Lemma DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large := WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large.wf . Lemma wf : well_founded WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large) . clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2_large. Qed. End WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2. Definition wf_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2 := WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2.wf . Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2) . intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))). apply wf_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2. Qed. Lemma wf : well_founded WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large . Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_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____ (x13:Z) (x14:Z) := 2 + 3* x13 + 1* x14. Definition P_id_a := 0. Definition P_id_a__U42 (x13:Z) (x14:Z) := 0. Definition P_id_U42 (x13:Z) (x14:Z) := 0. Definition P_id_a__U21 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_U21 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_a__U72 (x13:Z) := 0. Definition P_id_U72 (x13:Z) := 0. Definition P_id_a__U11 (x13:Z) (x14:Z) := 0. Definition P_id_u := 2. Definition P_id_a__U53 (x13:Z) := 0. Definition P_id_U53 (x13:Z) := 0. Definition P_id_a__U31 (x13:Z) (x14:Z) := 0. Definition P_id_U31 (x13:Z) (x14:Z) := 0. Definition P_id_isPalListKind (x13:Z) := 2* x13. Definition P_id_mark (x13:Z) := 1* x13. Definition P_id_i := 0. Definition P_id_a__U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_a__isList (x13:Z) := 0. Definition P_id_isList (x13:Z) := 0. Definition P_id_a__and (x13:Z) (x14:Z) := 2* x13 + 1* x14. Definition P_id_a__U12 (x13:Z) := 0. Definition P_id_U12 (x13:Z) := 0. Definition P_id_a__U62 (x13:Z) := 0. Definition P_id_U62 (x13:Z) := 0. Definition P_id_a__isQid (x13:Z) := 0. Definition P_id_isQid (x13:Z) := 0. Definition P_id_isPal (x13:Z) := 0. Definition P_id___ (x13:Z) (x14:Z) := 2 + 3* x13 + 1* x14. Definition P_id_e := 2. Definition P_id_a__U43 (x13:Z) := 2* x13. Definition P_id_U43 (x13:Z) := 2* x13. Definition P_id_a__U22 (x13:Z) (x14:Z) := 0. Definition P_id_U22 (x13:Z) (x14:Z) := 0. Definition P_id_a__isNePal (x13:Z) := 1* x13. Definition P_id_isNePal (x13:Z) := 1* x13. Definition P_id_tt := 0. Definition P_id_U11 (x13:Z) (x14:Z) := 0. Definition P_id_a__U61 (x13:Z) (x14:Z) := 0. Definition P_id_U61 (x13:Z) (x14:Z) := 0. Definition P_id_a__U32 (x13:Z) := 0. Definition P_id_U32 (x13:Z) := 0. Definition P_id_and (x13:Z) (x14:Z) := 2* x13 + 1* x14. Definition P_id_nil := 0. Definition P_id_o := 0. Definition P_id_a__U52 (x13:Z) (x14:Z) := 0. Definition P_id_U52 (x13:Z) (x14:Z) := 0. Definition P_id_a__U23 (x13:Z) := 0. Definition P_id_U23 (x13:Z) := 0. Definition P_id_a__isPalListKind (x13:Z) := 2* x13. Definition P_id_a__isNeList (x13:Z) := 0. Definition P_id_isNeList (x13:Z) := 0. Definition P_id_a__U71 (x13:Z) (x14:Z) := 0. Definition P_id_U71 (x13:Z) (x14:Z) := 0. Definition P_id_a__U41 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_U41 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_a__isPal (x13:Z) := 0. Lemma P_id_a_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a____ x14 x16 <= P_id_a____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U42 x14 x16 <= P_id_a__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U42 x14 x16 <= P_id_U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U21 x14 x16 x18 <= P_id_a__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U21 x14 x16 x18 <= P_id_U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U72 x14 <= P_id_a__U72 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U72 x14 <= P_id_U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U11 x14 x16 <= P_id_a__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U53 x14 <= P_id_a__U53 x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U53 x14 <= P_id_U53 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U31 x14 x16 <= P_id_a__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U31 x14 x16 <= P_id_U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isPalListKind x14 <= P_id_isPalListKind x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_mark x14 <= P_id_mark x13. Proof. intros x14 x13. intros [H_1 H_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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U51 x14 x16 x18 <= P_id_a__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U51_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U51 x14 x16 x18 <= P_id_U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isList x14 <= P_id_a__isList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isList x14 <= P_id_isList x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__and x14 x16 <= P_id_a__and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U12_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U12 x14 <= P_id_a__U12 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U12 x14 <= P_id_U12 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U62 x14 <= P_id_a__U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U62 x14 <= P_id_U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isQid x14 <= P_id_a__isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isQid x14 <= P_id_isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isPal x14 <= P_id_isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id___ x14 x16 <= P_id___ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U43_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U43 x14 <= P_id_a__U43 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U43 x14 <= P_id_U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U22 x14 x16 <= P_id_a__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U22 x14 x16 <= P_id_U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNePal x14 <= P_id_a__isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isNePal x14 <= P_id_isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U11_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U11 x14 x16 <= P_id_U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U61 x14 x16 <= P_id_a__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U61 x14 x16 <= P_id_U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U32 x14 <= P_id_a__U32 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U32 x14 <= P_id_U32 x13. Proof. intros x14 x13. intros [H_1 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_and x14 x16 <= P_id_and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U52 x14 x16 <= P_id_a__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U52 x14 x16 <= P_id_U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U23 x14 <= P_id_a__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U23 x14 <= P_id_U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPalListKind x14 <= P_id_a__isPalListKind x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNeList x14 <= P_id_a__isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isNeList x14 <= P_id_isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U71 x14 x16 <= P_id_a__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U71 x14 x16 <= P_id_U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U41 x14 x16 x18 <= P_id_a__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U41_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U41 x14 x16 x18 <= P_id_U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPal x14 <= P_id_a__isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a____ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U42 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U42 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U21 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U21 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U72 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_U72 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U11 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_u_bounded : 0 <= P_id_u . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U53 x13. 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 x13, (0 <= x13) ->0 <= P_id_U53 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U31 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U31 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPalListKind x13. 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 x13, (0 <= x13) ->0 <= P_id_mark x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_i_bounded : 0 <= P_id_i . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_bounded : forall x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U51 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U51 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isList x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__and x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U12 x13. 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 x13, (0 <= x13) ->0 <= P_id_U12 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id___ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_e_bounded : 0 <= P_id_e . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U43 x13. 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 x13, (0 <= x13) ->0 <= P_id_U43 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U22 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U22 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNePal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNePal x13. 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_U11_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U11 x14 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U61 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U61 x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U32 x13. 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 x13, (0 <= x13) ->0 <= P_id_U32 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_and x14 x13. 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_o_bounded : 0 <= P_id_o . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U52 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U52 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPalListKind x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U71 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U71 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U41 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U41 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPal x13. 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____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_a____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) => P_id_a__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U42 (x14:: x13::nil)) => P_id_U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) => P_id_a__U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U21 (x15:: x14::x13::nil)) => P_id_U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_a__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_U72 (x13::nil)) => P_id_U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) => P_id_a__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_a__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_U53 (x13::nil)) => P_id_U53 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) => P_id_a__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U31 (x14:: x13::nil)) => P_id_U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil)) => P_id_isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_mark (measure x13) | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) => P_id_a__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U51 (x15:: x14::x13::nil)) => P_id_U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_a__isList (measure x13) | (algebra.Alg.Term algebra.F.id_isList (x13::nil)) => P_id_isList (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) => P_id_a__and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_a__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_U12 (x13::nil)) => P_id_U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_a__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_U62 (x13::nil)) => P_id_U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_a__isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isQid (x13::nil)) => P_id_isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isPal (x13::nil)) => P_id_isPal (measure x13) | (algebra.Alg.Term algebra.F.id___ (x14:: x13::nil)) => P_id___ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_a__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_U43 (x13::nil)) => P_id_U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) => P_id_a__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U22 (x14:: x13::nil)) => P_id_U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_a__isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_isNePal (x13::nil)) => P_id_isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U11 (x14:: x13::nil)) => P_id_U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) => P_id_a__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U61 (x14:: x13::nil)) => P_id_U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_a__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_U32 (x13::nil)) => P_id_U32 (measure x13) | (algebra.Alg.Term algebra.F.id_and (x14:: x13::nil)) => P_id_and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o | (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) => P_id_a__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U52 (x14:: x13::nil)) => P_id_U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_a__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_U23 (x13::nil)) => P_id_U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_a__isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_a__isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_isNeList (x13::nil)) => P_id_isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) => P_id_a__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U71 (x14:: x13::nil)) => P_id_U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) => P_id_a__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U41 (x15:: x14::x13::nil)) => P_id_U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_a__isPal (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__U22 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNEPAL (x13:Z) := 0. Definition P_id_A__U43 (x13:Z) := 0. Definition P_id_A__U32 (x13:Z) := 0. Definition P_id_A__U61 (x13:Z) (x14:Z) := 0. Definition P_id_A__U11 (x13:Z) (x14:Z) := 0. Definition P_id_A__U23 (x13:Z) := 0. Definition P_id_A__ISPALLISTKIND (x13:Z) := 2* x13. Definition P_id_A__U52 (x13:Z) (x14:Z) := 0. Definition P_id_A____ (x13:Z) (x14:Z) := 0. Definition P_id_A__U41 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_A__U71 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNELIST (x13:Z) := 0. Definition P_id_A__ISLIST (x13:Z) := 0. Definition P_id_A__AND (x13:Z) (x14:Z) := 1* x14. Definition P_id_A__U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_A__ISQID (x13:Z) := 0. Definition P_id_A__U62 (x13:Z) := 0. Definition P_id_A__U12 (x13:Z) := 0. Definition P_id_A__U31 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISPAL (x13:Z) := 0. Definition P_id_A__U53 (x13:Z) := 0. Definition P_id_MARK (x13:Z) := 1* x13. Definition P_id_A__U42 (x13:Z) (x14:Z) := 0. Definition P_id_A__U72 (x13:Z) := 0. Definition P_id_A__U21 (x13:Z) (x14:Z) (x15:Z) := 0. Lemma P_id_A__U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U22 x14 x16 <= P_id_A__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNEPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNEPAL x14 <= P_id_A__ISNEPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U43 x14 <= P_id_A__U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U32 x14 <= P_id_A__U32 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U61 x14 x16 <= P_id_A__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U11 x14 x16 <= P_id_A__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U23 x14 <= P_id_A__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISPALLISTKIND_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPALLISTKIND x14 <= P_id_A__ISPALLISTKIND x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U52 x14 x16 <= P_id_A__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A____ x14 x16 <= P_id_A____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U41 x14 x16 x18 <= P_id_A__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U71 x14 x16 <= P_id_A__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNELIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNELIST x14 <= P_id_A__ISNELIST x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISLIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISLIST x14 <= P_id_A__ISLIST x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__AND x14 x16 <= P_id_A__AND x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U51 x14 x16 x18 <= P_id_A__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__ISQID_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISQID x14 <= P_id_A__ISQID x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ; intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U62 x14 <= P_id_A__U62 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U12 x14 <= P_id_A__U12 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U31 x14 x16 <= P_id_A__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPAL x14 <= P_id_A__ISPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U53 x14 <= P_id_A__U53 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_MARK x14 <= P_id_MARK x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U42 x14 x16 <= P_id_A__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U72 x14 <= P_id_A__U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U21 x14 x16 x18 <= P_id_A__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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. Definition marked_measure := InterpZ.marked_measure 0 P_id_a____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal P_id_A__U22 P_id_A__ISNEPAL P_id_A__U43 P_id_A__U32 P_id_A__U61 P_id_A__U11 P_id_A__U23 P_id_A__ISPALLISTKIND P_id_A__U52 P_id_A____ P_id_A__U41 P_id_A__U71 P_id_A__ISNELIST P_id_A__ISLIST P_id_A__AND P_id_A__U51 P_id_A__ISQID P_id_A__U62 P_id_A__U12 P_id_A__U31 P_id_A__ISPAL P_id_A__U53 P_id_MARK P_id_A__U42 P_id_A__U72 P_id_A__U21. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_A__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_A__ISNEPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_A__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_A__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_A__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_A__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_A__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_A__ISPALLISTKIND (measure x13) | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_A__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) => P_id_A____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) => P_id_A__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_A__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_A__ISNELIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_A__ISLIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_A__AND (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14:: x13::nil)) => P_id_A__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_A__ISQID (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_A__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_A__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_A__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_A__ISPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_A__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_MARK (measure x13) | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_A__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_A__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14:: x13::nil)) => P_id_A__U21 (measure x15) (measure x14) (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__ISNEPAL_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U23_monotonic;assumption. intros ;apply P_id_A__ISPALLISTKIND_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A_____monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__U71_monotonic;assumption. intros ;apply P_id_A__ISNELIST_monotonic;assumption. intros ;apply P_id_A__ISLIST_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__ISQID_monotonic;assumption. intros ;apply P_id_A__U62_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__ISPAL_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__U72_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b). Definition le a b := marked_measure a <= marked_measure b. Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c. Proof. unfold lt, le in *. intros a b c. apply (interp.le_lt_compat_right (interp.o_Z 0)). Qed. Lemma wf_lt : well_founded lt. Proof. unfold lt in *. apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. Qed. Lemma DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large := WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large.wf . Lemma wf : well_founded WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large) . clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_large. Qed. End WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_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____ (x13:Z) (x14:Z) := 2 + 1* x13 + 1* x14. Definition P_id_a := 0. Definition P_id_a__U42 (x13:Z) (x14:Z) := 1 + 2* x14. Definition P_id_U42 (x13:Z) (x14:Z) := 1 + 2* x14. Definition P_id_a__U21 (x13:Z) (x14:Z) (x15:Z) := 2 + 1* x15. Definition P_id_U21 (x13:Z) (x14:Z) (x15:Z) := 2 + 1* x15. Definition P_id_a__U72 (x13:Z) := 0. Definition P_id_U72 (x13:Z) := 0. Definition P_id_a__U11 (x13:Z) (x14:Z) := 0. Definition P_id_u := 0. Definition P_id_a__U53 (x13:Z) := 0. Definition P_id_U53 (x13:Z) := 0. Definition P_id_a__U31 (x13:Z) (x14:Z) := 2* x13 + 2* x14. Definition P_id_U31 (x13:Z) (x14:Z) := 2* x13 + 2* x14. Definition P_id_isPalListKind (x13:Z) := 0. Definition P_id_mark (x13:Z) := 1* x13. Definition P_id_i := 0. Definition P_id_a__U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_a__isList (x13:Z) := 1* x13. Definition P_id_isList (x13:Z) := 1* x13. Definition P_id_a__and (x13:Z) (x14:Z) := 1* x13 + 2* x14. Definition P_id_a__U12 (x13:Z) := 0. Definition P_id_U12 (x13:Z) := 0. Definition P_id_a__U62 (x13:Z) := 0. Definition P_id_U62 (x13:Z) := 0. Definition P_id_a__isQid (x13:Z) := 0. Definition P_id_isQid (x13:Z) := 0. Definition P_id_isPal (x13:Z) := 0. Definition P_id___ (x13:Z) (x14:Z) := 2 + 1* x13 + 1* x14. Definition P_id_e := 0. Definition P_id_a__U43 (x13:Z) := 1* x13. Definition P_id_U43 (x13:Z) := 1* x13. Definition P_id_a__U22 (x13:Z) (x14:Z) := 2 + 1* x14. Definition P_id_U22 (x13:Z) (x14:Z) := 2 + 1* x14. Definition P_id_a__isNePal (x13:Z) := 0. Definition P_id_isNePal (x13:Z) := 0. Definition P_id_tt := 0. Definition P_id_U11 (x13:Z) (x14:Z) := 0. Definition P_id_a__U61 (x13:Z) (x14:Z) := 2* x13. Definition P_id_U61 (x13:Z) (x14:Z) := 2* x13. Definition P_id_a__U32 (x13:Z) := 0. Definition P_id_U32 (x13:Z) := 0. Definition P_id_and (x13:Z) (x14:Z) := 1* x13 + 2* x14. Definition P_id_nil := 2. Definition P_id_o := 0. Definition P_id_a__U52 (x13:Z) (x14:Z) := 0. Definition P_id_U52 (x13:Z) (x14:Z) := 0. Definition P_id_a__U23 (x13:Z) := 2 + 1* x13. Definition P_id_U23 (x13:Z) := 2 + 1* x13. Definition P_id_a__isPalListKind (x13:Z) := 0. Definition P_id_a__isNeList (x13:Z) := 2* x13. Definition P_id_isNeList (x13:Z) := 2* x13. Definition P_id_a__U71 (x13:Z) (x14:Z) := 0. Definition P_id_U71 (x13:Z) (x14:Z) := 0. Definition P_id_a__U41 (x13:Z) (x14:Z) (x15:Z) := 2 + 1* x14 + 2* x15. Definition P_id_U41 (x13:Z) (x14:Z) (x15:Z) := 2 + 1* x14 + 2* x15. Definition P_id_a__isPal (x13:Z) := 0. Lemma P_id_a_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a____ x14 x16 <= P_id_a____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U42 x14 x16 <= P_id_a__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U42 x14 x16 <= P_id_U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U21 x14 x16 x18 <= P_id_a__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U21 x14 x16 x18 <= P_id_U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U72 x14 <= P_id_a__U72 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U72 x14 <= P_id_U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U11 x14 x16 <= P_id_a__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U53 x14 <= P_id_a__U53 x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U53 x14 <= P_id_U53 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U31 x14 x16 <= P_id_a__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U31 x14 x16 <= P_id_U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isPalListKind x14 <= P_id_isPalListKind x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_mark x14 <= P_id_mark x13. Proof. intros x14 x13. intros [H_1 H_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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U51 x14 x16 x18 <= P_id_a__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U51_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U51 x14 x16 x18 <= P_id_U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isList x14 <= P_id_a__isList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isList x14 <= P_id_isList x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__and x14 x16 <= P_id_a__and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U12_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U12 x14 <= P_id_a__U12 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U12 x14 <= P_id_U12 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U62 x14 <= P_id_a__U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U62 x14 <= P_id_U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isQid x14 <= P_id_a__isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isQid x14 <= P_id_isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isPal x14 <= P_id_isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id___ x14 x16 <= P_id___ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U43_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U43 x14 <= P_id_a__U43 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U43 x14 <= P_id_U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U22 x14 x16 <= P_id_a__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U22 x14 x16 <= P_id_U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNePal x14 <= P_id_a__isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isNePal x14 <= P_id_isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U11_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U11 x14 x16 <= P_id_U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U61 x14 x16 <= P_id_a__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U61 x14 x16 <= P_id_U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U32 x14 <= P_id_a__U32 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U32 x14 <= P_id_U32 x13. Proof. intros x14 x13. intros [H_1 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_and x14 x16 <= P_id_and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U52 x14 x16 <= P_id_a__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U52 x14 x16 <= P_id_U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U23 x14 <= P_id_a__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U23 x14 <= P_id_U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPalListKind x14 <= P_id_a__isPalListKind x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNeList x14 <= P_id_a__isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isNeList x14 <= P_id_isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U71 x14 x16 <= P_id_a__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U71 x14 x16 <= P_id_U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U41 x14 x16 x18 <= P_id_a__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U41_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U41 x14 x16 x18 <= P_id_U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPal x14 <= P_id_a__isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a____ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U42 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U42 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U21 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U21 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U72 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_U72 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U11 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_u_bounded : 0 <= P_id_u . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U53 x13. 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 x13, (0 <= x13) ->0 <= P_id_U53 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U31 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U31 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPalListKind x13. 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 x13, (0 <= x13) ->0 <= P_id_mark x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_i_bounded : 0 <= P_id_i . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_bounded : forall x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U51 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U51 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isList x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__and x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U12 x13. 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 x13, (0 <= x13) ->0 <= P_id_U12 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id___ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_e_bounded : 0 <= P_id_e . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U43 x13. 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 x13, (0 <= x13) ->0 <= P_id_U43 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U22 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U22 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNePal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNePal x13. 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_U11_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U11 x14 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U61 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U61 x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U32 x13. 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 x13, (0 <= x13) ->0 <= P_id_U32 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_and x14 x13. 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_o_bounded : 0 <= P_id_o . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U52 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U52 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPalListKind x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U71 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U71 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U41 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U41 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPal x13. 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____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_a____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_a__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U42 (x14:: x13::nil)) => P_id_U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15:: x14::x13::nil)) => P_id_a__U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U21 (x15:: x14::x13::nil)) => P_id_U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_a__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_U72 (x13::nil)) => P_id_U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_a__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_a__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_U53 (x13::nil)) => P_id_U53 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_a__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U31 (x14:: x13::nil)) => P_id_U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil)) => P_id_isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_mark (measure x13) | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i | (algebra.Alg.Term algebra.F.id_a__U51 (x15:: x14::x13::nil)) => P_id_a__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U51 (x15:: x14::x13::nil)) => P_id_U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_a__isList (measure x13) | (algebra.Alg.Term algebra.F.id_isList (x13::nil)) => P_id_isList (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_a__and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_a__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_U12 (x13::nil)) => P_id_U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_a__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_U62 (x13::nil)) => P_id_U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_a__isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isQid (x13::nil)) => P_id_isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isPal (x13::nil)) => P_id_isPal (measure x13) | (algebra.Alg.Term algebra.F.id___ (x14:: x13::nil)) => P_id___ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_a__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_U43 (x13::nil)) => P_id_U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_a__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U22 (x14:: x13::nil)) => P_id_U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_a__isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_isNePal (x13::nil)) => P_id_isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U11 (x14:: x13::nil)) => P_id_U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_a__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U61 (x14:: x13::nil)) => P_id_U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_a__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_U32 (x13::nil)) => P_id_U32 (measure x13) | (algebra.Alg.Term algebra.F.id_and (x14:: x13::nil)) => P_id_and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_a__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U52 (x14:: x13::nil)) => P_id_U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_a__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_U23 (x13::nil)) => P_id_U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_a__isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_a__isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_isNeList (x13::nil)) => P_id_isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_a__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U71 (x14:: x13::nil)) => P_id_U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15:: x14::x13::nil)) => P_id_a__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U41 (x15:: x14::x13::nil)) => P_id_U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_a__isPal (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__U22 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNEPAL (x13:Z) := 0. Definition P_id_A__U43 (x13:Z) := 0. Definition P_id_A__U32 (x13:Z) := 0. Definition P_id_A__U61 (x13:Z) (x14:Z) := 0. Definition P_id_A__U11 (x13:Z) (x14:Z) := 0. Definition P_id_A__U23 (x13:Z) := 0. Definition P_id_A__ISPALLISTKIND (x13:Z) := 0. Definition P_id_A__U52 (x13:Z) (x14:Z) := 0. Definition P_id_A____ (x13:Z) (x14:Z) := 0. Definition P_id_A__U41 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_A__U71 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNELIST (x13:Z) := 0. Definition P_id_A__ISLIST (x13:Z) := 0. Definition P_id_A__AND (x13:Z) (x14:Z) := 2* x14. Definition P_id_A__U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_A__ISQID (x13:Z) := 0. Definition P_id_A__U62 (x13:Z) := 0. Definition P_id_A__U12 (x13:Z) := 0. Definition P_id_A__U31 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISPAL (x13:Z) := 0. Definition P_id_A__U53 (x13:Z) := 0. Definition P_id_MARK (x13:Z) := 2* x13. Definition P_id_A__U42 (x13:Z) (x14:Z) := 0. Definition P_id_A__U72 (x13:Z) := 0. Definition P_id_A__U21 (x13:Z) (x14:Z) (x15:Z) := 0. Lemma P_id_A__U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U22 x14 x16 <= P_id_A__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNEPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNEPAL x14 <= P_id_A__ISNEPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U43 x14 <= P_id_A__U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U32 x14 <= P_id_A__U32 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U61 x14 x16 <= P_id_A__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U11 x14 x16 <= P_id_A__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U23 x14 <= P_id_A__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISPALLISTKIND_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPALLISTKIND x14 <= P_id_A__ISPALLISTKIND x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U52 x14 x16 <= P_id_A__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A____ x14 x16 <= P_id_A____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U41 x14 x16 x18 <= P_id_A__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U71 x14 x16 <= P_id_A__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNELIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNELIST x14 <= P_id_A__ISNELIST x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISLIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISLIST x14 <= P_id_A__ISLIST x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__AND x14 x16 <= P_id_A__AND x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U51 x14 x16 x18 <= P_id_A__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__ISQID_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISQID x14 <= P_id_A__ISQID x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U62 x14 <= P_id_A__U62 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U12 x14 <= P_id_A__U12 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U31 x14 x16 <= P_id_A__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPAL x14 <= P_id_A__ISPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U53 x14 <= P_id_A__U53 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_MARK x14 <= P_id_MARK x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U42 x14 x16 <= P_id_A__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U72 x14 <= P_id_A__U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U21 x14 x16 x18 <= P_id_A__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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. Definition marked_measure := InterpZ.marked_measure 0 P_id_a____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal P_id_A__U22 P_id_A__ISNEPAL P_id_A__U43 P_id_A__U32 P_id_A__U61 P_id_A__U11 P_id_A__U23 P_id_A__ISPALLISTKIND P_id_A__U52 P_id_A____ P_id_A__U41 P_id_A__U71 P_id_A__ISNELIST P_id_A__ISLIST P_id_A__AND P_id_A__U51 P_id_A__ISQID P_id_A__U62 P_id_A__U12 P_id_A__U31 P_id_A__ISPAL P_id_A__U53 P_id_MARK P_id_A__U42 P_id_A__U72 P_id_A__U21. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) => P_id_A__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_A__ISNEPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_A__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_A__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) => P_id_A__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) => P_id_A__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_A__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_A__ISPALLISTKIND (measure x13) | (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) => P_id_A__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) => P_id_A____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) => P_id_A__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) => P_id_A__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_A__ISNELIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_A__ISLIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) => P_id_A__AND (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14:: x13::nil)) => P_id_A__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_A__ISQID (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_A__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_A__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) => P_id_A__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_A__ISPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_A__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_MARK (measure x13) | (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) => P_id_A__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_A__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14:: x13::nil)) => P_id_A__U21 (measure x15) (measure x14) (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__ISNEPAL_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U23_monotonic;assumption. intros ;apply P_id_A__ISPALLISTKIND_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A_____monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__U71_monotonic;assumption. intros ;apply P_id_A__ISNELIST_monotonic;assumption. intros ;apply P_id_A__ISLIST_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__ISQID_monotonic;assumption. intros ;apply P_id_A__U62_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__ISPAL_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__U72_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b). Definition le a b := marked_measure a <= marked_measure b. Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c. Proof. unfold lt, le in *. intros a b c. apply (interp.le_lt_compat_right (interp.o_Z 0)). Qed. Lemma wf_lt : well_founded lt. Proof. unfold lt in *. apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. Qed. Lemma DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large := WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large.wf . Lemma wf : well_founded WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_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_23_large_large_large_large_large_large_scc_2_large_scc_7_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_23_large_large_large_large_large_large_scc_2_large_scc_7_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_large. Qed. End WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_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____ (x13:Z) (x14:Z) := 2* x13 + 1* x14. Definition P_id_a := 0. Definition P_id_a__U42 (x13:Z) (x14:Z) := 2. Definition P_id_U42 (x13:Z) (x14:Z) := 2. Definition P_id_a__U21 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_U21 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_a__U72 (x13:Z) := 0. Definition P_id_U72 (x13:Z) := 0. Definition P_id_a__U11 (x13:Z) (x14:Z) := 0. Definition P_id_u := 0. Definition P_id_a__U53 (x13:Z) := 1 + 1* x13. Definition P_id_U53 (x13:Z) := 1 + 1* x13. Definition P_id_a__U31 (x13:Z) (x14:Z) := 0. Definition P_id_U31 (x13:Z) (x14:Z) := 0. Definition P_id_isPalListKind (x13:Z) := 0. Definition P_id_mark (x13:Z) := 1* x13. Definition P_id_i := 0. Definition P_id_a__U51 (x13:Z) (x14:Z) (x15:Z) := 1 + 2* x13. Definition P_id_U51 (x13:Z) (x14:Z) (x15:Z) := 1 + 2* x13. Definition P_id_a__isList (x13:Z) := 0. Definition P_id_isList (x13:Z) := 0. Definition P_id_a__and (x13:Z) (x14:Z) := 2* x13 + 2* x14. Definition P_id_a__U12 (x13:Z) := 0. Definition P_id_U12 (x13:Z) := 0. Definition P_id_a__U62 (x13:Z) := 0. Definition P_id_U62 (x13:Z) := 0. Definition P_id_a__isQid (x13:Z) := 0. Definition P_id_isQid (x13:Z) := 0. Definition P_id_isPal (x13:Z) := 0. Definition P_id___ (x13:Z) (x14:Z) := 2* x13 + 1* x14. Definition P_id_e := 0. Definition P_id_a__U43 (x13:Z) := 1* x13. Definition P_id_U43 (x13:Z) := 1* x13. Definition P_id_a__U22 (x13:Z) (x14:Z) := 1* x13. Definition P_id_U22 (x13:Z) (x14:Z) := 1* x13. Definition P_id_a__isNePal (x13:Z) := 0. Definition P_id_isNePal (x13:Z) := 0. Definition P_id_tt := 0. Definition P_id_U11 (x13:Z) (x14:Z) := 0. Definition P_id_a__U61 (x13:Z) (x14:Z) := 0. Definition P_id_U61 (x13:Z) (x14:Z) := 0. Definition P_id_a__U32 (x13:Z) := 2* x13. Definition P_id_U32 (x13:Z) := 2* x13. Definition P_id_and (x13:Z) (x14:Z) := 2* x13 + 2* x14. Definition P_id_nil := 1. Definition P_id_o := 0. Definition P_id_a__U52 (x13:Z) (x14:Z) := 1. Definition P_id_U52 (x13:Z) (x14:Z) := 1. Definition P_id_a__U23 (x13:Z) := 2* x13. Definition P_id_U23 (x13:Z) := 2* x13. Definition P_id_a__isPalListKind (x13:Z) := 0. Definition P_id_a__isNeList (x13:Z) := 2. Definition P_id_isNeList (x13:Z) := 2. Definition P_id_a__U71 (x13:Z) (x14:Z) := 0. Definition P_id_U71 (x13:Z) (x14:Z) := 0. Definition P_id_a__U41 (x13:Z) (x14:Z) (x15:Z) := 2. Definition P_id_U41 (x13:Z) (x14:Z) (x15:Z) := 2. Definition P_id_a__isPal (x13:Z) := 0. Lemma P_id_a_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a____ x14 x16 <= P_id_a____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U42 x14 x16 <= P_id_a__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U42 x14 x16 <= P_id_U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U21 x14 x16 x18 <= P_id_a__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U21 x14 x16 x18 <= P_id_U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U72 x14 <= P_id_a__U72 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U72 x14 <= P_id_U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U11 x14 x16 <= P_id_a__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U53 x14 <= P_id_a__U53 x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U53 x14 <= P_id_U53 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U31 x14 x16 <= P_id_a__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U31 x14 x16 <= P_id_U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isPalListKind x14 <= P_id_isPalListKind x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_mark x14 <= P_id_mark x13. Proof. intros x14 x13. intros [H_1 H_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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U51 x14 x16 x18 <= P_id_a__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U51_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U51 x14 x16 x18 <= P_id_U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isList x14 <= P_id_a__isList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isList x14 <= P_id_isList x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__and x14 x16 <= P_id_a__and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U12_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U12 x14 <= P_id_a__U12 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U12 x14 <= P_id_U12 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U62 x14 <= P_id_a__U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U62 x14 <= P_id_U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isQid x14 <= P_id_a__isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isQid x14 <= P_id_isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isPal x14 <= P_id_isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id___ x14 x16 <= P_id___ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U43_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U43 x14 <= P_id_a__U43 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U43 x14 <= P_id_U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U22 x14 x16 <= P_id_a__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U22 x14 x16 <= P_id_U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNePal x14 <= P_id_a__isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isNePal x14 <= P_id_isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U11_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U11 x14 x16 <= P_id_U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U61 x14 x16 <= P_id_a__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U61 x14 x16 <= P_id_U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U32 x14 <= P_id_a__U32 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U32 x14 <= P_id_U32 x13. Proof. intros x14 x13. intros [H_1 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_and x14 x16 <= P_id_and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U52 x14 x16 <= P_id_a__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U52 x14 x16 <= P_id_U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__U23 x14 <= P_id_a__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U23 x14 <= P_id_U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPalListKind x14 <= P_id_a__isPalListKind x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNeList x14 <= P_id_a__isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isNeList x14 <= P_id_isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U71 x14 x16 <= P_id_a__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U71 x14 x16 <= P_id_U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U41 x14 x16 x18 <= P_id_a__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U41_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U41 x14 x16 x18 <= P_id_U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPal x14 <= P_id_a__isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a____ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U42 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U42 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U21 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U21 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U72 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_U72 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U11 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_u_bounded : 0 <= P_id_u . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U53 x13. 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 x13, (0 <= x13) ->0 <= P_id_U53 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U31 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U31 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPalListKind x13. 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 x13, (0 <= x13) ->0 <= P_id_mark x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_i_bounded : 0 <= P_id_i . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_bounded : forall x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U51 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U51 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isList x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__and x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U12 x13. 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 x13, (0 <= x13) ->0 <= P_id_U12 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id___ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_e_bounded : 0 <= P_id_e . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U43 x13. 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 x13, (0 <= x13) ->0 <= P_id_U43 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U22 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U22 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNePal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNePal x13. 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_U11_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U11 x14 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U61 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U61 x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U32 x13. 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 x13, (0 <= x13) ->0 <= P_id_U32 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_and x14 x13. 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_o_bounded : 0 <= P_id_o . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U52 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U52 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPalListKind x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U71 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U71 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U41 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U41 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPal x13. 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____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_a____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_a__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U42 (x14:: x13::nil)) => P_id_U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15:: x14::x13::nil)) => P_id_a__U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U21 (x15:: x14::x13::nil)) => P_id_U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_a__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_U72 (x13::nil)) => P_id_U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_a__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_a__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_U53 (x13::nil)) => P_id_U53 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_a__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U31 (x14:: x13::nil)) => P_id_U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil)) => P_id_isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_mark (measure x13) | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i | (algebra.Alg.Term algebra.F.id_a__U51 (x15:: x14::x13::nil)) => P_id_a__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U51 (x15:: x14::x13::nil)) => P_id_U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_a__isList (measure x13) | (algebra.Alg.Term algebra.F.id_isList (x13::nil)) => P_id_isList (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_a__and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_a__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_U12 (x13::nil)) => P_id_U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_a__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_U62 (x13::nil)) => P_id_U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_a__isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isQid (x13::nil)) => P_id_isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isPal (x13::nil)) => P_id_isPal (measure x13) | (algebra.Alg.Term algebra.F.id___ (x14:: x13::nil)) => P_id___ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_a__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_U43 (x13::nil)) => P_id_U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_a__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U22 (x14:: x13::nil)) => P_id_U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_a__isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_isNePal (x13::nil)) => P_id_isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U11 (x14:: x13::nil)) => P_id_U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_a__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U61 (x14:: x13::nil)) => P_id_U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_a__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_U32 (x13::nil)) => P_id_U32 (measure x13) | (algebra.Alg.Term algebra.F.id_and (x14:: x13::nil)) => P_id_and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_a__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U52 (x14:: x13::nil)) => P_id_U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_a__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_U23 (x13::nil)) => P_id_U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_a__isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_a__isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_isNeList (x13::nil)) => P_id_isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_a__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U71 (x14:: x13::nil)) => P_id_U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15:: x14::x13::nil)) => P_id_a__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U41 (x15:: x14::x13::nil)) => P_id_U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_a__isPal (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__U22 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNEPAL (x13:Z) := 0. Definition P_id_A__U43 (x13:Z) := 0. Definition P_id_A__U32 (x13:Z) := 0. Definition P_id_A__U61 (x13:Z) (x14:Z) := 0. Definition P_id_A__U11 (x13:Z) (x14:Z) := 0. Definition P_id_A__U23 (x13:Z) := 0. Definition P_id_A__ISPALLISTKIND (x13:Z) := 0. Definition P_id_A__U52 (x13:Z) (x14:Z) := 0. Definition P_id_A____ (x13:Z) (x14:Z) := 0. Definition P_id_A__U41 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_A__U71 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNELIST (x13:Z) := 0. Definition P_id_A__ISLIST (x13:Z) := 0. Definition P_id_A__AND (x13:Z) (x14:Z) := 1* x14. Definition P_id_A__U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_A__ISQID (x13:Z) := 0. Definition P_id_A__U62 (x13:Z) := 0. Definition P_id_A__U12 (x13:Z) := 0. Definition P_id_A__U31 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISPAL (x13:Z) := 0. Definition P_id_A__U53 (x13:Z) := 0. Definition P_id_MARK (x13:Z) := 1* x13. Definition P_id_A__U42 (x13:Z) (x14:Z) := 0. Definition P_id_A__U72 (x13:Z) := 0. Definition P_id_A__U21 (x13:Z) (x14:Z) (x15:Z) := 0. Lemma P_id_A__U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U22 x14 x16 <= P_id_A__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNEPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNEPAL x14 <= P_id_A__ISNEPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U43 x14 <= P_id_A__U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U32 x14 <= P_id_A__U32 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U61 x14 x16 <= P_id_A__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U11 x14 x16 <= P_id_A__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U23 x14 <= P_id_A__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISPALLISTKIND_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPALLISTKIND x14 <= P_id_A__ISPALLISTKIND x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U52 x14 x16 <= P_id_A__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A____ x14 x16 <= P_id_A____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U41 x14 x16 x18 <= P_id_A__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U71 x14 x16 <= P_id_A__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNELIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNELIST x14 <= P_id_A__ISNELIST x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISLIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISLIST x14 <= P_id_A__ISLIST x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__AND x14 x16 <= P_id_A__AND x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U51 x14 x16 x18 <= P_id_A__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__ISQID_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISQID x14 <= P_id_A__ISQID x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U62 x14 <= P_id_A__U62 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U12 x14 <= P_id_A__U12 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U31 x14 x16 <= P_id_A__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPAL x14 <= P_id_A__ISPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U53 x14 <= P_id_A__U53 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_MARK x14 <= P_id_MARK x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U42 x14 x16 <= P_id_A__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__U72 x14 <= P_id_A__U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U21 x14 x16 x18 <= P_id_A__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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. Definition marked_measure := InterpZ.marked_measure 0 P_id_a____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal P_id_A__U22 P_id_A__ISNEPAL P_id_A__U43 P_id_A__U32 P_id_A__U61 P_id_A__U11 P_id_A__U23 P_id_A__ISPALLISTKIND P_id_A__U52 P_id_A____ P_id_A__U41 P_id_A__U71 P_id_A__ISNELIST P_id_A__ISLIST P_id_A__AND P_id_A__U51 P_id_A__ISQID P_id_A__U62 P_id_A__U12 P_id_A__U31 P_id_A__ISPAL P_id_A__U53 P_id_MARK P_id_A__U42 P_id_A__U72 P_id_A__U21. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) => P_id_A__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_A__ISNEPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_A__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_A__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) => P_id_A__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) => P_id_A__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_A__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_A__ISPALLISTKIND (measure x13) | (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) => P_id_A__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) => P_id_A____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) => P_id_A__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) => P_id_A__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_A__ISNELIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_A__ISLIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) => P_id_A__AND (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) => P_id_A__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_A__ISQID (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_A__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_A__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) => P_id_A__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_A__ISPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_A__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_MARK (measure x13) | (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) => P_id_A__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_A__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) => P_id_A__U21 (measure x15) (measure x14) (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__ISNEPAL_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U23_monotonic;assumption. intros ;apply P_id_A__ISPALLISTKIND_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A_____monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__U71_monotonic;assumption. intros ;apply P_id_A__ISNELIST_monotonic;assumption. intros ;apply P_id_A__ISLIST_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__ISQID_monotonic;assumption. intros ;apply P_id_A__U62_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__ISPAL_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__U72_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b). Definition le a b := marked_measure a <= marked_measure b. Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c. Proof. unfold lt, le in *. intros a b c. apply (interp.le_lt_compat_right (interp.o_Z 0)). Qed. Lemma wf_lt : well_founded lt. Proof. unfold lt in *. apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. Qed. Lemma DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large := WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large.wf . Lemma wf : well_founded WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large) . clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4_large. Qed. End WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4. Definition wf_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4 := WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4.wf . Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). apply wf_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4. Qed. Lemma wf : well_founded WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large . Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_4; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_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____ (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_a := 0. Definition P_id_a__U42 (x13:Z) (x14:Z) := 1 + 1* x13. Definition P_id_U42 (x13:Z) (x14:Z) := 1 + 1* x13. Definition P_id_a__U21 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_U21 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_a__U72 (x13:Z) := 0. Definition P_id_U72 (x13:Z) := 0. Definition P_id_a__U11 (x13:Z) (x14:Z) := 0. Definition P_id_u := 1. Definition P_id_a__U53 (x13:Z) := 2* x13. Definition P_id_U53 (x13:Z) := 2* x13. Definition P_id_a__U31 (x13:Z) (x14:Z) := 1 + 1* x13. Definition P_id_U31 (x13:Z) (x14:Z) := 1 + 1* x13. Definition P_id_isPalListKind (x13:Z) := 0. Definition P_id_mark (x13:Z) := 1* x13. Definition P_id_i := 0. Definition P_id_a__U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_a__isList (x13:Z) := 0. Definition P_id_isList (x13:Z) := 0. Definition P_id_a__and (x13:Z) (x14:Z) := 1* x13 + 2* x14. Definition P_id_a__U12 (x13:Z) := 0. Definition P_id_U12 (x13:Z) := 0. Definition P_id_a__U62 (x13:Z) := 0. Definition P_id_U62 (x13:Z) := 0. Definition P_id_a__isQid (x13:Z) := 0. Definition P_id_isQid (x13:Z) := 0. Definition P_id_isPal (x13:Z) := 0. Definition P_id___ (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_e := 0. Definition P_id_a__U43 (x13:Z) := 1* x13. Definition P_id_U43 (x13:Z) := 1* x13. Definition P_id_a__U22 (x13:Z) (x14:Z) := 0. Definition P_id_U22 (x13:Z) (x14:Z) := 0. Definition P_id_a__isNePal (x13:Z) := 0. Definition P_id_isNePal (x13:Z) := 0. Definition P_id_tt := 0. Definition P_id_U11 (x13:Z) (x14:Z) := 0. Definition P_id_a__U61 (x13:Z) (x14:Z) := 1* x13. Definition P_id_U61 (x13:Z) (x14:Z) := 1* x13. Definition P_id_a__U32 (x13:Z) := 1 + 1* x13. Definition P_id_U32 (x13:Z) := 1 + 1* x13. Definition P_id_and (x13:Z) (x14:Z) := 1* x13 + 2* x14. Definition P_id_nil := 0. Definition P_id_o := 0. Definition P_id_a__U52 (x13:Z) (x14:Z) := 0. Definition P_id_U52 (x13:Z) (x14:Z) := 0. Definition P_id_a__U23 (x13:Z) := 2* x13. Definition P_id_U23 (x13:Z) := 2* x13. Definition P_id_a__isPalListKind (x13:Z) := 0. Definition P_id_a__isNeList (x13:Z) := 1. Definition P_id_isNeList (x13:Z) := 1. Definition P_id_a__U71 (x13:Z) (x14:Z) := 0. Definition P_id_U71 (x13:Z) (x14:Z) := 0. Definition P_id_a__U41 (x13:Z) (x14:Z) (x15:Z) := 1. Definition P_id_U41 (x13:Z) (x14:Z) (x15:Z) := 1. Definition P_id_a__isPal (x13:Z) := 0. Lemma P_id_a_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a____ x14 x16 <= P_id_a____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U42 x14 x16 <= P_id_a__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U42 x14 x16 <= P_id_U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U21 x14 x16 x18 <= P_id_a__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U21 x14 x16 x18 <= P_id_U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U72 x14 <= P_id_a__U72 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U72 x14 <= P_id_U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U11 x14 x16 <= P_id_a__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U53 x14 <= P_id_a__U53 x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U53 x14 <= P_id_U53 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U31 x14 x16 <= P_id_a__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U31 x14 x16 <= P_id_U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isPalListKind x14 <= P_id_isPalListKind x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_mark x14 <= P_id_mark x13. Proof. intros x14 x13. intros [H_1 H_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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U51 x14 x16 x18 <= P_id_a__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U51_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U51 x14 x16 x18 <= P_id_U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isList x14 <= P_id_a__isList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isList x14 <= P_id_isList x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__and x14 x16 <= P_id_a__and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U12_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U12 x14 <= P_id_a__U12 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U12 x14 <= P_id_U12 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U62 x14 <= P_id_a__U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U62 x14 <= P_id_U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isQid x14 <= P_id_a__isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isQid x14 <= P_id_isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isPal x14 <= P_id_isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id___ x14 x16 <= P_id___ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U43_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U43 x14 <= P_id_a__U43 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U43 x14 <= P_id_U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U22 x14 x16 <= P_id_a__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U22 x14 x16 <= P_id_U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNePal x14 <= P_id_a__isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isNePal x14 <= P_id_isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U11_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U11 x14 x16 <= P_id_U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U61 x14 x16 <= P_id_a__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U61 x14 x16 <= P_id_U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U32 x14 <= P_id_a__U32 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U32 x14 <= P_id_U32 x13. Proof. intros x14 x13. intros [H_1 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_and x14 x16 <= P_id_and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U52 x14 x16 <= P_id_a__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U52 x14 x16 <= P_id_U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U23 x14 <= P_id_a__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U23 x14 <= P_id_U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPalListKind x14 <= P_id_a__isPalListKind x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNeList x14 <= P_id_a__isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isNeList x14 <= P_id_isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U71 x14 x16 <= P_id_a__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U71 x14 x16 <= P_id_U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U41 x14 x16 x18 <= P_id_a__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U41_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U41 x14 x16 x18 <= P_id_U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPal x14 <= P_id_a__isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a____ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U42 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U42 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U21 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U21 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U72 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_U72 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U11 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_u_bounded : 0 <= P_id_u . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U53 x13. 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 x13, (0 <= x13) ->0 <= P_id_U53 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U31 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U31 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPalListKind x13. 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 x13, (0 <= x13) ->0 <= P_id_mark x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_i_bounded : 0 <= P_id_i . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_bounded : forall x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U51 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U51 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isList x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__and x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U12 x13. 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 x13, (0 <= x13) ->0 <= P_id_U12 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id___ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_e_bounded : 0 <= P_id_e . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U43 x13. 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 x13, (0 <= x13) ->0 <= P_id_U43 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U22 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U22 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNePal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNePal x13. 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_U11_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U11 x14 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U61 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U61 x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U32 x13. 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 x13, (0 <= x13) ->0 <= P_id_U32 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_and x14 x13. 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_o_bounded : 0 <= P_id_o . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U52 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U52 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPalListKind x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U71 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U71 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U41 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_U41 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPal x13. 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____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_a____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_a__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U42 (x14:: x13::nil)) => P_id_U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15:: x14::x13::nil)) => P_id_a__U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U21 (x15::x14:: x13::nil)) => P_id_U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_a__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_U72 (x13::nil)) => P_id_U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_a__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_a__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_U53 (x13::nil)) => P_id_U53 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_a__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U31 (x14:: x13::nil)) => P_id_U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil)) => P_id_isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_mark (measure x13) | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i | (algebra.Alg.Term algebra.F.id_a__U51 (x15:: x14::x13::nil)) => P_id_a__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U51 (x15::x14:: x13::nil)) => P_id_U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_a__isList (measure x13) | (algebra.Alg.Term algebra.F.id_isList (x13::nil)) => P_id_isList (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_a__and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_a__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_U12 (x13::nil)) => P_id_U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_a__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_U62 (x13::nil)) => P_id_U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_a__isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isQid (x13::nil)) => P_id_isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isPal (x13::nil)) => P_id_isPal (measure x13) | (algebra.Alg.Term algebra.F.id___ (x14:: x13::nil)) => P_id___ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_a__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_U43 (x13::nil)) => P_id_U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_a__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U22 (x14:: x13::nil)) => P_id_U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_a__isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_isNePal (x13::nil)) => P_id_isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U11 (x14:: x13::nil)) => P_id_U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_a__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U61 (x14:: x13::nil)) => P_id_U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_a__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_U32 (x13::nil)) => P_id_U32 (measure x13) | (algebra.Alg.Term algebra.F.id_and (x14:: x13::nil)) => P_id_and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_a__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U52 (x14:: x13::nil)) => P_id_U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_a__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_U23 (x13::nil)) => P_id_U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_a__isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_a__isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_isNeList (x13::nil)) => P_id_isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_a__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U71 (x14:: x13::nil)) => P_id_U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15:: x14::x13::nil)) => P_id_a__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U41 (x15::x14:: x13::nil)) => P_id_U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_a__isPal (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__U22 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNEPAL (x13:Z) := 0. Definition P_id_A__U43 (x13:Z) := 0. Definition P_id_A__U32 (x13:Z) := 0. Definition P_id_A__U61 (x13:Z) (x14:Z) := 0. Definition P_id_A__U11 (x13:Z) (x14:Z) := 0. Definition P_id_A__U23 (x13:Z) := 0. Definition P_id_A__ISPALLISTKIND (x13:Z) := 0. Definition P_id_A__U52 (x13:Z) (x14:Z) := 0. Definition P_id_A____ (x13:Z) (x14:Z) := 0. Definition P_id_A__U41 (x13:Z) (x14:Z) (x15:Z) := 2. Definition P_id_A__U71 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNELIST (x13:Z) := 1. Definition P_id_A__ISLIST (x13:Z) := 0. Definition P_id_A__AND (x13:Z) (x14:Z) := 2* x14. Definition P_id_A__U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_A__ISQID (x13:Z) := 0. Definition P_id_A__U62 (x13:Z) := 0. Definition P_id_A__U12 (x13:Z) := 0. Definition P_id_A__U31 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISPAL (x13:Z) := 0. Definition P_id_A__U53 (x13:Z) := 0. Definition P_id_MARK (x13:Z) := 2* x13. Definition P_id_A__U42 (x13:Z) (x14:Z) := 2. Definition P_id_A__U72 (x13:Z) := 0. Definition P_id_A__U21 (x13:Z) (x14:Z) (x15:Z) := 0. Lemma P_id_A__U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U22 x14 x16 <= P_id_A__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNEPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNEPAL x14 <= P_id_A__ISNEPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U43 x14 <= P_id_A__U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U32 x14 <= P_id_A__U32 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U61 x14 x16 <= P_id_A__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U11 x14 x16 <= P_id_A__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U23 x14 <= P_id_A__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISPALLISTKIND_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPALLISTKIND x14 <= P_id_A__ISPALLISTKIND x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U52 x14 x16 <= P_id_A__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A____ x14 x16 <= P_id_A____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U41 x14 x16 x18 <= P_id_A__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U71 x14 x16 <= P_id_A__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNELIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNELIST x14 <= P_id_A__ISNELIST x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISLIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISLIST x14 <= P_id_A__ISLIST x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__AND x14 x16 <= P_id_A__AND x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U51 x14 x16 x18 <= P_id_A__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__ISQID_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISQID x14 <= P_id_A__ISQID x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U62 x14 <= P_id_A__U62 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U12 x14 <= P_id_A__U12 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U31 x14 x16 <= P_id_A__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPAL x14 <= P_id_A__ISPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U53 x14 <= P_id_A__U53 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_MARK x14 <= P_id_MARK x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U42 x14 x16 <= P_id_A__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U72 x14 <= P_id_A__U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U21 x14 x16 x18 <= P_id_A__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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. Definition marked_measure := InterpZ.marked_measure 0 P_id_a____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal P_id_A__U22 P_id_A__ISNEPAL P_id_A__U43 P_id_A__U32 P_id_A__U61 P_id_A__U11 P_id_A__U23 P_id_A__ISPALLISTKIND P_id_A__U52 P_id_A____ P_id_A__U41 P_id_A__U71 P_id_A__ISNELIST P_id_A__ISLIST P_id_A__AND P_id_A__U51 P_id_A__ISQID P_id_A__U62 P_id_A__U12 P_id_A__U31 P_id_A__ISPAL P_id_A__U53 P_id_MARK P_id_A__U42 P_id_A__U72 P_id_A__U21. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) => P_id_A__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_A__ISNEPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_A__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_A__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) => P_id_A__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) => P_id_A__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_A__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_A__ISPALLISTKIND (measure x13) | (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) => P_id_A__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) => P_id_A____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) => P_id_A__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) => P_id_A__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_A__ISNELIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_A__ISLIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) => P_id_A__AND (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) => P_id_A__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_A__ISQID (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_A__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_A__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) => P_id_A__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_A__ISPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_A__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_MARK (measure x13) | (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) => P_id_A__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_A__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) => P_id_A__U21 (measure x15) (measure x14) (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__ISNEPAL_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U23_monotonic;assumption. intros ;apply P_id_A__ISPALLISTKIND_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A_____monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__U71_monotonic;assumption. intros ;apply P_id_A__ISNELIST_monotonic;assumption. intros ;apply P_id_A__ISLIST_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__ISQID_monotonic;assumption. intros ;apply P_id_A__U62_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__ISPAL_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__U72_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b). Definition le a b := marked_measure a <= marked_measure b. Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c. Proof. unfold lt, le in *. intros a b c. apply (interp.le_lt_compat_right (interp.o_Z 0)). Qed. Lemma wf_lt : well_founded lt. Proof. unfold lt in *. apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. Qed. Lemma DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large := WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large.wf . Lemma wf : well_founded WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large) . clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_large_scc_7_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7_large. Qed. End WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7. Definition wf_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 := WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7.wf . Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_large_non_scc_6; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_5; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_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_23_large_large_large_large_large_large_scc_2_large_scc_7. Qed. Lemma wf : well_founded WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2.DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large . Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_6; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_5; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_4; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_7; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_6; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_5; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_4; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (fail)))))))))))))))). Qed. End WF_DP_R_xml_0_scc_23_large_large_large_large_large_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____ (x13:Z) (x14:Z) := 1 + 2* x13 + 1* x14. Definition P_id_a := 0. Definition P_id_a__U42 (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_U42 (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_a__U21 (x13:Z) (x14:Z) (x15:Z) := 2 + 2* x14 + 1* x15. Definition P_id_U21 (x13:Z) (x14:Z) (x15:Z) := 2 + 2* x14 + 1* x15. Definition P_id_a__U72 (x13:Z) := 0. Definition P_id_U72 (x13:Z) := 0. Definition P_id_a__U11 (x13:Z) (x14:Z) := 1 + 1* x13 + 1* x14. Definition P_id_u := 0. Definition P_id_a__U53 (x13:Z) := 1* x13. Definition P_id_U53 (x13:Z) := 1* x13. Definition P_id_a__U31 (x13:Z) (x14:Z) := 1* x13. Definition P_id_U31 (x13:Z) (x14:Z) := 1* x13. Definition P_id_isPalListKind (x13:Z) := 0. Definition P_id_mark (x13:Z) := 1* x13. Definition P_id_i := 0. Definition P_id_a__U51 (x13:Z) (x14:Z) (x15:Z) := 1 + 1* x15. Definition P_id_U51 (x13:Z) (x14:Z) (x15:Z) := 1 + 1* x15. Definition P_id_a__isList (x13:Z) := 1 + 1* x13. Definition P_id_isList (x13:Z) := 1 + 1* x13. Definition P_id_a__and (x13:Z) (x14:Z) := 2* x13 + 2* x14. Definition P_id_a__U12 (x13:Z) := 1 + 1* x13. Definition P_id_U12 (x13:Z) := 1 + 1* x13. Definition P_id_a__U62 (x13:Z) := 0. Definition P_id_U62 (x13:Z) := 0. Definition P_id_a__isQid (x13:Z) := 0. Definition P_id_isQid (x13:Z) := 0. Definition P_id_isPal (x13:Z) := 0. Definition P_id___ (x13:Z) (x14:Z) := 1 + 2* x13 + 1* x14. Definition P_id_e := 0. Definition P_id_a__U43 (x13:Z) := 1* x13. Definition P_id_U43 (x13:Z) := 1* x13. Definition P_id_a__U22 (x13:Z) (x14:Z) := 1 + 1* x13 + 1* x14. Definition P_id_U22 (x13:Z) (x14:Z) := 1 + 1* x13 + 1* x14. Definition P_id_a__isNePal (x13:Z) := 0. Definition P_id_isNePal (x13:Z) := 0. Definition P_id_tt := 0. Definition P_id_U11 (x13:Z) (x14:Z) := 1 + 1* x13 + 1* x14. Definition P_id_a__U61 (x13:Z) (x14:Z) := 0. Definition P_id_U61 (x13:Z) (x14:Z) := 0. Definition P_id_a__U32 (x13:Z) := 1* x13. Definition P_id_U32 (x13:Z) := 1* x13. Definition P_id_and (x13:Z) (x14:Z) := 2* x13 + 2* x14. Definition P_id_nil := 0. Definition P_id_o := 2. Definition P_id_a__U52 (x13:Z) (x14:Z) := 1 + 1* x14. Definition P_id_U52 (x13:Z) (x14:Z) := 1 + 1* x14. Definition P_id_a__U23 (x13:Z) := 1* x13. Definition P_id_U23 (x13:Z) := 1* x13. Definition P_id_a__isPalListKind (x13:Z) := 0. Definition P_id_a__isNeList (x13:Z) := 1* x13. Definition P_id_isNeList (x13:Z) := 1* x13. Definition P_id_a__U71 (x13:Z) (x14:Z) := 0. Definition P_id_U71 (x13:Z) (x14:Z) := 0. Definition P_id_a__U41 (x13:Z) (x14:Z) (x15:Z) := 1 + 1* x13 + 2* x14 + 1* x15. Definition P_id_U41 (x13:Z) (x14:Z) (x15:Z) := 1 + 1* x13 + 2* x14 + 1* x15. Definition P_id_a__isPal (x13:Z) := 0. Lemma P_id_a_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a____ x14 x16 <= P_id_a____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U42 x14 x16 <= P_id_a__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U42 x14 x16 <= P_id_U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U21 x14 x16 x18 <= P_id_a__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U21 x14 x16 x18 <= P_id_U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U72 x14 <= P_id_a__U72 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U72 x14 <= P_id_U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U11 x14 x16 <= P_id_a__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U53 x14 <= P_id_a__U53 x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U53 x14 <= P_id_U53 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U31 x14 x16 <= P_id_a__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U31 x14 x16 <= P_id_U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isPalListKind x14 <= P_id_isPalListKind x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_mark x14 <= P_id_mark x13. Proof. intros x14 x13. intros [H_1 H_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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U51 x14 x16 x18 <= P_id_a__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U51_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U51 x14 x16 x18 <= P_id_U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isList x14 <= P_id_a__isList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isList x14 <= P_id_isList x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__and x14 x16 <= P_id_a__and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U12_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U12 x14 <= P_id_a__U12 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U12 x14 <= P_id_U12 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U62 x14 <= P_id_a__U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U62 x14 <= P_id_U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isQid x14 <= P_id_a__isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isQid x14 <= P_id_isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isPal x14 <= P_id_isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id___ x14 x16 <= P_id___ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U43_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U43 x14 <= P_id_a__U43 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U43 x14 <= P_id_U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U22 x14 x16 <= P_id_a__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U22 x14 x16 <= P_id_U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNePal x14 <= P_id_a__isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isNePal x14 <= P_id_isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U11_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U11 x14 x16 <= P_id_U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U61 x14 x16 <= P_id_a__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U61 x14 x16 <= P_id_U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U32 x14 <= P_id_a__U32 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U32 x14 <= P_id_U32 x13. Proof. intros x14 x13. intros [H_1 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_and x14 x16 <= P_id_and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U52 x14 x16 <= P_id_a__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U52 x14 x16 <= P_id_U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U23 x14 <= P_id_a__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U23 x14 <= P_id_U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPalListKind x14 <= P_id_a__isPalListKind x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNeList x14 <= P_id_a__isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isNeList x14 <= P_id_isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U71 x14 x16 <= P_id_a__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U71 x14 x16 <= P_id_U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U41 x14 x16 x18 <= P_id_a__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U41_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U41 x14 x16 x18 <= P_id_U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPal x14 <= P_id_a__isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a____ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U42 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U42 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U21 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U21 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U72 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_U72 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U11 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_u_bounded : 0 <= P_id_u . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U53 x13. 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 x13, (0 <= x13) ->0 <= P_id_U53 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U31 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U31 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPalListKind x13. 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 x13, (0 <= x13) ->0 <= P_id_mark x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_i_bounded : 0 <= P_id_i . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_bounded : forall x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U51 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U51 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isList x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__and x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U12 x13. 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 x13, (0 <= x13) ->0 <= P_id_U12 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id___ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_e_bounded : 0 <= P_id_e . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U43 x13. 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 x13, (0 <= x13) ->0 <= P_id_U43 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U22 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U22 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNePal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNePal x13. 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_U11_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U11 x14 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U61 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U61 x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U32 x13. 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 x13, (0 <= x13) ->0 <= P_id_U32 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_and x14 x13. 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_o_bounded : 0 <= P_id_o . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U52 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U52 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPalListKind x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U71 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U71 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) -> (0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U41 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U41 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPal x13. 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____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_a____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_a__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U42 (x14:: x13::nil)) => P_id_U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14:: x13::nil)) => P_id_a__U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U21 (x15::x14:: x13::nil)) => P_id_U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_a__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_U72 (x13::nil)) => P_id_U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_a__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_a__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_U53 (x13::nil)) => P_id_U53 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_a__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U31 (x14:: x13::nil)) => P_id_U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil)) => P_id_isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_mark (measure x13) | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14:: x13::nil)) => P_id_a__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U51 (x15::x14:: x13::nil)) => P_id_U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_a__isList (measure x13) | (algebra.Alg.Term algebra.F.id_isList (x13::nil)) => P_id_isList (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_a__and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_a__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_U12 (x13::nil)) => P_id_U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_a__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_U62 (x13::nil)) => P_id_U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_a__isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isQid (x13::nil)) => P_id_isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isPal (x13::nil)) => P_id_isPal (measure x13) | (algebra.Alg.Term algebra.F.id___ (x14:: x13::nil)) => P_id___ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_a__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_U43 (x13::nil)) => P_id_U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_a__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U22 (x14:: x13::nil)) => P_id_U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_a__isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_isNePal (x13::nil)) => P_id_isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U11 (x14:: x13::nil)) => P_id_U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_a__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U61 (x14:: x13::nil)) => P_id_U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_a__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_U32 (x13::nil)) => P_id_U32 (measure x13) | (algebra.Alg.Term algebra.F.id_and (x14:: x13::nil)) => P_id_and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_a__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U52 (x14:: x13::nil)) => P_id_U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_a__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_U23 (x13::nil)) => P_id_U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_a__isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_a__isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_isNeList (x13::nil)) => P_id_isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_a__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U71 (x14:: x13::nil)) => P_id_U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) => P_id_a__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U41 (x15::x14:: x13::nil)) => P_id_U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_a__isPal (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__U22 (x13:Z) (x14:Z) := 1 + 1* x14. Definition P_id_A__ISNEPAL (x13:Z) := 0. Definition P_id_A__U43 (x13:Z) := 0. Definition P_id_A__U32 (x13:Z) := 0. Definition P_id_A__U61 (x13:Z) (x14:Z) := 0. Definition P_id_A__U11 (x13:Z) (x14:Z) := 1 + 1* x14. Definition P_id_A__U23 (x13:Z) := 0. Definition P_id_A__ISPALLISTKIND (x13:Z) := 0. Definition P_id_A__U52 (x13:Z) (x14:Z) := 0. Definition P_id_A____ (x13:Z) (x14:Z) := 0. Definition P_id_A__U41 (x13:Z) (x14:Z) (x15:Z) := 1 + 2* x14 + 1* x15. Definition P_id_A__U71 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNELIST (x13:Z) := 1* x13. Definition P_id_A__ISLIST (x13:Z) := 1 + 1* x13. Definition P_id_A__AND (x13:Z) (x14:Z) := 2* x14. Definition P_id_A__U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_A__ISQID (x13:Z) := 0. Definition P_id_A__U62 (x13:Z) := 0. Definition P_id_A__U12 (x13:Z) := 0. Definition P_id_A__U31 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISPAL (x13:Z) := 0. Definition P_id_A__U53 (x13:Z) := 0. Definition P_id_MARK (x13:Z) := 1* x13. Definition P_id_A__U42 (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_A__U72 (x13:Z) := 0. Definition P_id_A__U21 (x13:Z) (x14:Z) (x15:Z) := 0. Lemma P_id_A__U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U22 x14 x16 <= P_id_A__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNEPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNEPAL x14 <= P_id_A__ISNEPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U43 x14 <= P_id_A__U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U32 x14 <= P_id_A__U32 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U61 x14 x16 <= P_id_A__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U11 x14 x16 <= P_id_A__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U23 x14 <= P_id_A__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISPALLISTKIND_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPALLISTKIND x14 <= P_id_A__ISPALLISTKIND x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U52 x14 x16 <= P_id_A__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A____ x14 x16 <= P_id_A____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U41 x14 x16 x18 <= P_id_A__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U71 x14 x16 <= P_id_A__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNELIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNELIST x14 <= P_id_A__ISNELIST x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISLIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISLIST x14 <= P_id_A__ISLIST x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__AND x14 x16 <= P_id_A__AND x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U51 x14 x16 x18 <= P_id_A__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__ISQID_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISQID x14 <= P_id_A__ISQID x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U62 x14 <= P_id_A__U62 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U12 x14 <= P_id_A__U12 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U31 x14 x16 <= P_id_A__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPAL x14 <= P_id_A__ISPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U53 x14 <= P_id_A__U53 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_MARK x14 <= P_id_MARK x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U42 x14 x16 <= P_id_A__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U72 x14 <= P_id_A__U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U21 x14 x16 x18 <= P_id_A__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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. Definition marked_measure := InterpZ.marked_measure 0 P_id_a____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal P_id_A__U22 P_id_A__ISNEPAL P_id_A__U43 P_id_A__U32 P_id_A__U61 P_id_A__U11 P_id_A__U23 P_id_A__ISPALLISTKIND P_id_A__U52 P_id_A____ P_id_A__U41 P_id_A__U71 P_id_A__ISNELIST P_id_A__ISLIST P_id_A__AND P_id_A__U51 P_id_A__ISQID P_id_A__U62 P_id_A__U12 P_id_A__U31 P_id_A__ISPAL P_id_A__U53 P_id_MARK P_id_A__U42 P_id_A__U72 P_id_A__U21. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) => P_id_A__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_A__ISNEPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_A__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_A__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) => P_id_A__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) => P_id_A__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_A__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_A__ISPALLISTKIND (measure x13) | (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) => P_id_A__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) => P_id_A____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14::x13::nil)) => P_id_A__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) => P_id_A__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_A__ISNELIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_A__ISLIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) => P_id_A__AND (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14::x13::nil)) => P_id_A__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_A__ISQID (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_A__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_A__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) => P_id_A__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_A__ISPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_A__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_MARK (measure x13) | (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) => P_id_A__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_A__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14::x13::nil)) => P_id_A__U21 (measure x15) (measure x14) (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__ISNEPAL_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U23_monotonic;assumption. intros ;apply P_id_A__ISPALLISTKIND_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A_____monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__U71_monotonic;assumption. intros ;apply P_id_A__ISNELIST_monotonic;assumption. intros ;apply P_id_A__ISLIST_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__ISQID_monotonic;assumption. intros ;apply P_id_A__U62_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__ISPAL_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__U72_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b). Definition le a b := marked_measure a <= marked_measure b. Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c. Proof. unfold lt, le in *. intros a b c. apply (interp.le_lt_compat_right (interp.o_Z 0)). Qed. Lemma wf_lt : well_founded lt. Proof. unfold lt in *. apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. Qed. Lemma DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large_large_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_23_large_large_large_large_large_large_scc_2_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large_large_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_23_large_large_large_large_large_large_scc_2_large := WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large.wf . Lemma wf : well_founded WF_DP_R_xml_0_scc_23_large_large_large_large_large_large.DP_R_xml_0_scc_23_large_large_large_large_large_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_23_large_large_large_large_large_large_scc_2_large) . clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_23_large_large_large_large_large_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_23_large_large_large_large_large_large_scc_2_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2_large. Qed. End WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2. Definition wf_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 := WF_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2.wf. Lemma acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 : forall x y, (DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2 x y) -> Acc WF_DP_R_xml_0_scc_23_large_large_large_large_large.DP_R_xml_0_scc_23_large_large_large_large_large_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_23_large_large_large_large_large_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_23_large_large_large_large_large_large_scc_2. Qed. Lemma wf : well_founded WF_DP_R_xml_0_scc_23_large_large_large_large_large.DP_R_xml_0_scc_23_large_large_large_large_large_large . Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_large_large_large_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23_large_large_large_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_23_large_large_large_large_large_large. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_a____ (x13:Z) (x14:Z) := 2 + 2* x13 + 1* x14. Definition P_id_a := 0. Definition P_id_a__U42 (x13:Z) (x14:Z) := 1* x13 + 2* x14. Definition P_id_U42 (x13:Z) (x14:Z) := 1* x13 + 2* x14. Definition P_id_a__U21 (x13:Z) (x14:Z) (x15:Z) := 1 + 2* x13 + 2* x14 + 2* x15. Definition P_id_U21 (x13:Z) (x14:Z) (x15:Z) := 1 + 2* x13 + 2* x14 + 2* x15. Definition P_id_a__U72 (x13:Z) := 0. Definition P_id_U72 (x13:Z) := 0. Definition P_id_a__U11 (x13:Z) (x14:Z) := 2* x13 + 2* x14. Definition P_id_u := 0. Definition P_id_a__U53 (x13:Z) := 1* x13. Definition P_id_U53 (x13:Z) := 1* x13. Definition P_id_a__U31 (x13:Z) (x14:Z) := 2* x13. Definition P_id_U31 (x13:Z) (x14:Z) := 2* x13. Definition P_id_isPalListKind (x13:Z) := 0. Definition P_id_mark (x13:Z) := 1* x13. Definition P_id_i := 0. Definition P_id_a__U51 (x13:Z) (x14:Z) (x15:Z) := 1 + 1* x13 + 2* x14 + 2* x15. Definition P_id_U51 (x13:Z) (x14:Z) (x15:Z) := 1 + 1* x13 + 2* x14 + 2* x15. Definition P_id_a__isList (x13:Z) := 2* x13. Definition P_id_isList (x13:Z) := 2* x13. Definition P_id_a__and (x13:Z) (x14:Z) := 2* x13 + 2* x14. Definition P_id_a__U12 (x13:Z) := 1* x13. Definition P_id_U12 (x13:Z) := 1* x13. Definition P_id_a__U62 (x13:Z) := 0. Definition P_id_U62 (x13:Z) := 0. Definition P_id_a__isQid (x13:Z) := 0. Definition P_id_isQid (x13:Z) := 0. Definition P_id_isPal (x13:Z) := 0. Definition P_id___ (x13:Z) (x14:Z) := 2 + 2* x13 + 1* x14. Definition P_id_e := 0. Definition P_id_a__U43 (x13:Z) := 1* x13. Definition P_id_U43 (x13:Z) := 1* x13. Definition P_id_a__U22 (x13:Z) (x14:Z) := 1* x13 + 2* x14. Definition P_id_U22 (x13:Z) (x14:Z) := 1* x13 + 2* x14. Definition P_id_a__isNePal (x13:Z) := 1. Definition P_id_isNePal (x13:Z) := 1. Definition P_id_tt := 0. Definition P_id_U11 (x13:Z) (x14:Z) := 2* x13 + 2* x14. Definition P_id_a__U61 (x13:Z) (x14:Z) := 1. Definition P_id_U61 (x13:Z) (x14:Z) := 1. Definition P_id_a__U32 (x13:Z) := 1* x13. Definition P_id_U32 (x13:Z) := 1* x13. Definition P_id_and (x13:Z) (x14:Z) := 2* x13 + 2* x14. Definition P_id_nil := 0. Definition P_id_o := 1. Definition P_id_a__U52 (x13:Z) (x14:Z) := 1 + 1* x13 + 2* x14. Definition P_id_U52 (x13:Z) (x14:Z) := 1 + 1* x13 + 2* x14. Definition P_id_a__U23 (x13:Z) := 1* x13. Definition P_id_U23 (x13:Z) := 1* x13. Definition P_id_a__isPalListKind (x13:Z) := 0. Definition P_id_a__isNeList (x13:Z) := 2* x13. Definition P_id_isNeList (x13:Z) := 2* x13. Definition P_id_a__U71 (x13:Z) (x14:Z) := 0. Definition P_id_U71 (x13:Z) (x14:Z) := 0. Definition P_id_a__U41 (x13:Z) (x14:Z) (x15:Z) := 1* x13 + 2* x14 + 2* x15. Definition P_id_U41 (x13:Z) (x14:Z) (x15:Z) := 1* x13 + 2* x14 + 2* x15. Definition P_id_a__isPal (x13:Z) := 0. Lemma P_id_a_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a____ x14 x16 <= P_id_a____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U42 x14 x16 <= P_id_a__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U42 x14 x16 <= P_id_U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U21 x14 x16 x18 <= P_id_a__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U21 x14 x16 x18 <= P_id_U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U72 x14 <= P_id_a__U72 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U72 x14 <= P_id_U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U11 x14 x16 <= P_id_a__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U53 x14 <= P_id_a__U53 x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U53 x14 <= P_id_U53 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U31 x14 x16 <= P_id_a__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U31 x14 x16 <= P_id_U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isPalListKind x14 <= P_id_isPalListKind x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_mark x14 <= P_id_mark x13. Proof. intros x14 x13. intros [H_1 H_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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U51 x14 x16 x18 <= P_id_a__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U51_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U51 x14 x16 x18 <= P_id_U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isList x14 <= P_id_a__isList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isList x14 <= P_id_isList x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__and x14 x16 <= P_id_a__and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U12_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U12 x14 <= P_id_a__U12 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U12 x14 <= P_id_U12 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U62 x14 <= P_id_a__U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U62 x14 <= P_id_U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isQid x14 <= P_id_a__isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isQid x14 <= P_id_isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isPal x14 <= P_id_isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id___ x14 x16 <= P_id___ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U43_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U43 x14 <= P_id_a__U43 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U43 x14 <= P_id_U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U22 x14 x16 <= P_id_a__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U22 x14 x16 <= P_id_U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNePal x14 <= P_id_a__isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isNePal x14 <= P_id_isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U11_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U11 x14 x16 <= P_id_U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U61 x14 x16 <= P_id_a__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U61 x14 x16 <= P_id_U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U32 x14 <= P_id_a__U32 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U32 x14 <= P_id_U32 x13. Proof. intros x14 x13. intros [H_1 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_and x14 x16 <= P_id_and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U52 x14 x16 <= P_id_a__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U52 x14 x16 <= P_id_U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U23 x14 <= P_id_a__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U23 x14 <= P_id_U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPalListKind x14 <= P_id_a__isPalListKind x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNeList x14 <= P_id_a__isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isNeList x14 <= P_id_isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U71 x14 x16 <= P_id_a__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U71 x14 x16 <= P_id_U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U41 x14 x16 x18 <= P_id_a__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U41_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U41 x14 x16 x18 <= P_id_U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isPal x14 <= P_id_a__isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a____ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U42 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U42 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U21 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U21 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U72 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_U72 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U11 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_u_bounded : 0 <= P_id_u . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U53 x13. 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 x13, (0 <= x13) ->0 <= P_id_U53 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U31 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U31 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPalListKind x13. 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 x13, (0 <= x13) ->0 <= P_id_mark x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_i_bounded : 0 <= P_id_i . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_bounded : forall x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U51 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U51 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isList x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__and x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U12 x13. 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 x13, (0 <= x13) ->0 <= P_id_U12 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id___ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_e_bounded : 0 <= P_id_e . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U43 x13. 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 x13, (0 <= x13) ->0 <= P_id_U43 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U22 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U22 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNePal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNePal x13. 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_U11_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U11 x14 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U61 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U61 x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U32 x13. 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 x13, (0 <= x13) ->0 <= P_id_U32 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_and x14 x13. 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_o_bounded : 0 <= P_id_o . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U52 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U52 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPalListKind x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U71 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U71 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U41 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U41 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPal x13. 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____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_a____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_a__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U42 (x14::x13::nil)) => P_id_U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14:: x13::nil)) => P_id_a__U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U21 (x15::x14:: x13::nil)) => P_id_U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_a__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_U72 (x13::nil)) => P_id_U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_a__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_a__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_U53 (x13::nil)) => P_id_U53 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_a__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U31 (x14::x13::nil)) => P_id_U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil)) => P_id_isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_mark (measure x13) | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14:: x13::nil)) => P_id_a__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U51 (x15::x14:: x13::nil)) => P_id_U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_a__isList (measure x13) | (algebra.Alg.Term algebra.F.id_isList (x13::nil)) => P_id_isList (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_a__and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_a__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_U12 (x13::nil)) => P_id_U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_a__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_U62 (x13::nil)) => P_id_U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_a__isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isQid (x13::nil)) => P_id_isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isPal (x13::nil)) => P_id_isPal (measure x13) | (algebra.Alg.Term algebra.F.id___ (x14::x13::nil)) => P_id___ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_a__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_U43 (x13::nil)) => P_id_U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_a__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U22 (x14::x13::nil)) => P_id_U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_a__isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_isNePal (x13::nil)) => P_id_isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U11 (x14::x13::nil)) => P_id_U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_a__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U61 (x14::x13::nil)) => P_id_U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_a__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_U32 (x13::nil)) => P_id_U32 (measure x13) | (algebra.Alg.Term algebra.F.id_and (x14::x13::nil)) => P_id_and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_a__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U52 (x14::x13::nil)) => P_id_U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_a__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_U23 (x13::nil)) => P_id_U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_a__isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_a__isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_isNeList (x13::nil)) => P_id_isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_a__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U71 (x14::x13::nil)) => P_id_U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) => P_id_a__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U41 (x15::x14:: x13::nil)) => P_id_U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_a__isPal (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__U22 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNEPAL (x13:Z) := 0. Definition P_id_A__U43 (x13:Z) := 0. Definition P_id_A__U32 (x13:Z) := 0. Definition P_id_A__U61 (x13:Z) (x14:Z) := 0. Definition P_id_A__U11 (x13:Z) (x14:Z) := 0. Definition P_id_A__U23 (x13:Z) := 0. Definition P_id_A__ISPALLISTKIND (x13:Z) := 0. Definition P_id_A__U52 (x13:Z) (x14:Z) := 1 + 2* x14. Definition P_id_A____ (x13:Z) (x14:Z) := 0. Definition P_id_A__U41 (x13:Z) (x14:Z) (x15:Z) := 2* x15. Definition P_id_A__U71 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNELIST (x13:Z) := 0. Definition P_id_A__ISLIST (x13:Z) := 0. Definition P_id_A__AND (x13:Z) (x14:Z) := 2* x14. Definition P_id_A__U51 (x13:Z) (x14:Z) (x15:Z) := 1 + 3* x14 + 3* x15 . Definition P_id_A__ISQID (x13:Z) := 0. Definition P_id_A__U62 (x13:Z) := 0. Definition P_id_A__U12 (x13:Z) := 0. Definition P_id_A__U31 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISPAL (x13:Z) := 0. Definition P_id_A__U53 (x13:Z) := 0. Definition P_id_MARK (x13:Z) := 2* x13. Definition P_id_A__U42 (x13:Z) (x14:Z) := 1* x14. Definition P_id_A__U72 (x13:Z) := 0. Definition P_id_A__U21 (x13:Z) (x14:Z) (x15:Z) := 1 + 3* x15. Lemma P_id_A__U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U22 x14 x16 <= P_id_A__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNEPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNEPAL x14 <= P_id_A__ISNEPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U43 x14 <= P_id_A__U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U32 x14 <= P_id_A__U32 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U61 x14 x16 <= P_id_A__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U11 x14 x16 <= P_id_A__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U23 x14 <= P_id_A__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISPALLISTKIND_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPALLISTKIND x14 <= P_id_A__ISPALLISTKIND x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U52 x14 x16 <= P_id_A__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A____ x14 x16 <= P_id_A____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U41 x14 x16 x18 <= P_id_A__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U71 x14 x16 <= P_id_A__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNELIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNELIST x14 <= P_id_A__ISNELIST x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISLIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISLIST x14 <= P_id_A__ISLIST x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__AND x14 x16 <= P_id_A__AND x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U51 x14 x16 x18 <= P_id_A__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__ISQID_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISQID x14 <= P_id_A__ISQID x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U62 x14 <= P_id_A__U62 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U12 x14 <= P_id_A__U12 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U31 x14 x16 <= P_id_A__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISPAL x14 <= P_id_A__ISPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U53 x14 <= P_id_A__U53 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_MARK x14 <= P_id_MARK x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U42 x14 x16 <= P_id_A__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U72 x14 <= P_id_A__U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U21 x14 x16 x18 <= P_id_A__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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. Definition marked_measure := InterpZ.marked_measure 0 P_id_a____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal P_id_A__U22 P_id_A__ISNEPAL P_id_A__U43 P_id_A__U32 P_id_A__U61 P_id_A__U11 P_id_A__U23 P_id_A__ISPALLISTKIND P_id_A__U52 P_id_A____ P_id_A__U41 P_id_A__U71 P_id_A__ISNELIST P_id_A__ISLIST P_id_A__AND P_id_A__U51 P_id_A__ISQID P_id_A__U62 P_id_A__U12 P_id_A__U31 P_id_A__ISPAL P_id_A__U53 P_id_MARK P_id_A__U42 P_id_A__U72 P_id_A__U21. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_A__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_A__ISNEPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_A__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_A__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_A__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_A__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_A__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_A__ISPALLISTKIND (measure x13) | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_A__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_A____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15:: x14::x13::nil)) => P_id_A__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_A__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_A__ISNELIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_A__ISLIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_A__AND (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U51 (x15:: x14::x13::nil)) => P_id_A__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_A__ISQID (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_A__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_A__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_A__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_A__ISPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_A__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_MARK (measure x13) | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_A__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_A__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15:: x14::x13::nil)) => P_id_A__U21 (measure x15) (measure x14) (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__ISNEPAL_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U23_monotonic;assumption. intros ;apply P_id_A__ISPALLISTKIND_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A_____monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__U71_monotonic;assumption. intros ;apply P_id_A__ISNELIST_monotonic;assumption. intros ;apply P_id_A__ISLIST_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__ISQID_monotonic;assumption. intros ;apply P_id_A__U62_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__ISPAL_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__U72_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b). Definition le a b := marked_measure a <= marked_measure b. Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c. Proof. unfold lt, le in *. intros a b c. apply (interp.le_lt_compat_right (interp.o_Z 0)). Qed. Lemma wf_lt : well_founded lt. Proof. unfold lt in *. apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. Qed. Lemma DP_R_xml_0_scc_23_large_large_large_large_large_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large_large_strict lt. Proof. unfold Relation_Definitions.inclusion, lt in *. intros a b H;destruct H; match goal with | |- (Zwf.Zwf 0) _ (marked_measure (algebra.Alg.Term ?f ?l)) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_lt_compat_right (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos );(assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma DP_R_xml_0_scc_23_large_large_large_large_large_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large_large_large le. Proof. unfold Relation_Definitions.inclusion, le, Zwf.Zwf in *. intros a b H;destruct H; match goal with | |- _ <= marked_measure (algebra.Alg.Term ?f ?l) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_trans (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos );(assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition wf_DP_R_xml_0_scc_23_large_large_large_large_large_large := WF_DP_R_xml_0_scc_23_large_large_large_large_large_large.wf. Lemma wf : well_founded WF_DP_R_xml_0_scc_23_large_large_large_large.DP_R_xml_0_scc_23_large_large_large_large_large . Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_23_large_large_large_large_large_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_23_large_large_large_large_large_strict_in_lt; econstructor eassumption)|| ((apply IHx;[econstructor eassumption| intros y' Hlt;apply IHx';apply lt_le_compat with (1:=Hlt) ; apply DP_R_xml_0_scc_23_large_large_large_large_large_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_23_large_large_large_large_large_large. Qed. End WF_DP_R_xml_0_scc_23_large_large_large_large_large. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_a____ (x13:Z) (x14:Z) := 1 + 2* x13 + 1* x14. Definition P_id_a := 0. Definition P_id_a__U42 (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_U42 (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_a__U21 (x13:Z) (x14:Z) (x15:Z) := 1* x13 + 1* x14 + 1* x15. Definition P_id_U21 (x13:Z) (x14:Z) (x15:Z) := 1* x13 + 1* x14 + 1* x15. Definition P_id_a__U72 (x13:Z) := 0. Definition P_id_U72 (x13:Z) := 0. Definition P_id_a__U11 (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_u := 0. Definition P_id_a__U53 (x13:Z) := 1* x13. Definition P_id_U53 (x13:Z) := 1* x13. Definition P_id_a__U31 (x13:Z) (x14:Z) := 1* x13. Definition P_id_U31 (x13:Z) (x14:Z) := 1* x13. Definition P_id_isPalListKind (x13:Z) := 0. Definition P_id_mark (x13:Z) := 1* x13. Definition P_id_i := 0. Definition P_id_a__U51 (x13:Z) (x14:Z) (x15:Z) := 2* x13 + 2* x14 + 1* x15. Definition P_id_U51 (x13:Z) (x14:Z) (x15:Z) := 2* x13 + 2* x14 + 1* x15. Definition P_id_a__isList (x13:Z) := 1* x13. Definition P_id_isList (x13:Z) := 1* x13. Definition P_id_a__and (x13:Z) (x14:Z) := 2* x13 + 2* x14. Definition P_id_a__U12 (x13:Z) := 1* x13. Definition P_id_U12 (x13:Z) := 1* x13. Definition P_id_a__U62 (x13:Z) := 0. Definition P_id_U62 (x13:Z) := 0. Definition P_id_a__isQid (x13:Z) := 0. Definition P_id_isQid (x13:Z) := 0. Definition P_id_isPal (x13:Z) := 0. Definition P_id___ (x13:Z) (x14:Z) := 1 + 2* x13 + 1* x14. Definition P_id_e := 0. Definition P_id_a__U43 (x13:Z) := 1* x13. Definition P_id_U43 (x13:Z) := 1* x13. Definition P_id_a__U22 (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_U22 (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_a__isNePal (x13:Z) := 0. Definition P_id_isNePal (x13:Z) := 0. Definition P_id_tt := 0. Definition P_id_U11 (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_a__U61 (x13:Z) (x14:Z) := 0. Definition P_id_U61 (x13:Z) (x14:Z) := 0. Definition P_id_a__U32 (x13:Z) := 2* x13. Definition P_id_U32 (x13:Z) := 2* x13. Definition P_id_and (x13:Z) (x14:Z) := 2* x13 + 2* x14. Definition P_id_nil := 0. Definition P_id_o := 1. Definition P_id_a__U52 (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_U52 (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_a__U23 (x13:Z) := 1* x13. Definition P_id_U23 (x13:Z) := 1* x13. Definition P_id_a__isPalListKind (x13:Z) := 0. Definition P_id_a__isNeList (x13:Z) := 1* x13. Definition P_id_isNeList (x13:Z) := 1* x13. Definition P_id_a__U71 (x13:Z) (x14:Z) := 0. Definition P_id_U71 (x13:Z) (x14:Z) := 0. Definition P_id_a__U41 (x13:Z) (x14:Z) (x15:Z) := 2* x13 + 1* x14 + 1* x15. Definition P_id_U41 (x13:Z) (x14:Z) (x15:Z) := 2* x13 + 1* x14 + 1* x15. Definition P_id_a__isPal (x13:Z) := 0. Lemma P_id_a_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a____ x14 x16 <= P_id_a____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U42 x14 x16 <= P_id_a__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U42 x14 x16 <= P_id_U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U21 x14 x16 x18 <= P_id_a__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U21 x14 x16 x18 <= P_id_U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U72 x14 <= P_id_a__U72 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U72 x14 <= P_id_U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U11 x14 x16 <= P_id_a__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U53 x14 <= P_id_a__U53 x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U53 x14 <= P_id_U53 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U31 x14 x16 <= P_id_a__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U31 x14 x16 <= P_id_U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isPalListKind x14 <= P_id_isPalListKind x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_mark x14 <= P_id_mark x13. Proof. intros x14 x13. intros [H_1 H_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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U51 x14 x16 x18 <= P_id_a__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U51_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U51 x14 x16 x18 <= P_id_U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isList x14 <= P_id_a__isList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isList x14 <= P_id_isList x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__and x14 x16 <= P_id_a__and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U12_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U12 x14 <= P_id_a__U12 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U12 x14 <= P_id_U12 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U62 x14 <= P_id_a__U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U62 x14 <= P_id_U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isQid x14 <= P_id_a__isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isQid x14 <= P_id_isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isPal x14 <= P_id_isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id___ x14 x16 <= P_id___ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U43_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U43 x14 <= P_id_a__U43 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U43 x14 <= P_id_U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U22 x14 x16 <= P_id_a__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U22 x14 x16 <= P_id_U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNePal x14 <= P_id_a__isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isNePal x14 <= P_id_isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U11_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U11 x14 x16 <= P_id_U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U61 x14 x16 <= P_id_a__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U61 x14 x16 <= P_id_U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U32 x14 <= P_id_a__U32 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U32 x14 <= P_id_U32 x13. Proof. intros x14 x13. intros [H_1 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_and x14 x16 <= P_id_and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U52 x14 x16 <= P_id_a__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U52 x14 x16 <= P_id_U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U23 x14 <= P_id_a__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U23 x14 <= P_id_U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPalListKind x14 <= P_id_a__isPalListKind x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNeList x14 <= P_id_a__isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isNeList x14 <= P_id_isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U71 x14 x16 <= P_id_a__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U71 x14 x16 <= P_id_U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U41 x14 x16 x18 <= P_id_a__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U41_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U41 x14 x16 x18 <= P_id_U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isPal x14 <= P_id_a__isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a____ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U42 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U42 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U21 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U21 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U72 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_U72 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U11 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_u_bounded : 0 <= P_id_u . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U53 x13. 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 x13, (0 <= x13) ->0 <= P_id_U53 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U31 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U31 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPalListKind x13. 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 x13, (0 <= x13) ->0 <= P_id_mark x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_i_bounded : 0 <= P_id_i . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_bounded : forall x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U51 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U51 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isList x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__and x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U12 x13. 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 x13, (0 <= x13) ->0 <= P_id_U12 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id___ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_e_bounded : 0 <= P_id_e . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U43 x13. 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 x13, (0 <= x13) ->0 <= P_id_U43 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U22 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U22 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNePal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNePal x13. 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_U11_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U11 x14 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U61 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U61 x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U32 x13. 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 x13, (0 <= x13) ->0 <= P_id_U32 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_and x14 x13. 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_o_bounded : 0 <= P_id_o . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U52 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U52 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPalListKind x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U71 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U71 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U41 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U41 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPal x13. 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____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_a____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_a__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U42 (x14::x13::nil)) => P_id_U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14:: x13::nil)) => P_id_a__U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U21 (x15::x14:: x13::nil)) => P_id_U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_a__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_U72 (x13::nil)) => P_id_U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_a__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_a__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_U53 (x13::nil)) => P_id_U53 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_a__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U31 (x14::x13::nil)) => P_id_U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil)) => P_id_isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_mark (measure x13) | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14:: x13::nil)) => P_id_a__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U51 (x15::x14:: x13::nil)) => P_id_U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_a__isList (measure x13) | (algebra.Alg.Term algebra.F.id_isList (x13::nil)) => P_id_isList (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_a__and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_a__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_U12 (x13::nil)) => P_id_U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_a__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_U62 (x13::nil)) => P_id_U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_a__isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isQid (x13::nil)) => P_id_isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isPal (x13::nil)) => P_id_isPal (measure x13) | (algebra.Alg.Term algebra.F.id___ (x14::x13::nil)) => P_id___ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_a__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_U43 (x13::nil)) => P_id_U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_a__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U22 (x14::x13::nil)) => P_id_U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_a__isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_isNePal (x13::nil)) => P_id_isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U11 (x14::x13::nil)) => P_id_U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_a__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U61 (x14::x13::nil)) => P_id_U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_a__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_U32 (x13::nil)) => P_id_U32 (measure x13) | (algebra.Alg.Term algebra.F.id_and (x14::x13::nil)) => P_id_and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_a__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U52 (x14::x13::nil)) => P_id_U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_a__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_U23 (x13::nil)) => P_id_U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_a__isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_a__isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_isNeList (x13::nil)) => P_id_isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_a__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U71 (x14::x13::nil)) => P_id_U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) => P_id_a__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U41 (x15::x14:: x13::nil)) => P_id_U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_a__isPal (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__U22 (x13:Z) (x14:Z) := 1* x14. Definition P_id_A__ISNEPAL (x13:Z) := 0. Definition P_id_A__U43 (x13:Z) := 0. Definition P_id_A__U32 (x13:Z) := 0. Definition P_id_A__U61 (x13:Z) (x14:Z) := 0. Definition P_id_A__U11 (x13:Z) (x14:Z) := 1* x14. Definition P_id_A__U23 (x13:Z) := 0. Definition P_id_A__ISPALLISTKIND (x13:Z) := 0. Definition P_id_A__U52 (x13:Z) (x14:Z) := 1* x14. Definition P_id_A____ (x13:Z) (x14:Z) := 0. Definition P_id_A__U41 (x13:Z) (x14:Z) (x15:Z) := 2* x14 + 1* x15. Definition P_id_A__U71 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNELIST (x13:Z) := 1* x13. Definition P_id_A__ISLIST (x13:Z) := 1* x13. Definition P_id_A__AND (x13:Z) (x14:Z) := 2* x14. Definition P_id_A__U51 (x13:Z) (x14:Z) (x15:Z) := 2* x14 + 1* x15. Definition P_id_A__ISQID (x13:Z) := 0. Definition P_id_A__U62 (x13:Z) := 0. Definition P_id_A__U12 (x13:Z) := 0. Definition P_id_A__U31 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISPAL (x13:Z) := 0. Definition P_id_A__U53 (x13:Z) := 0. Definition P_id_MARK (x13:Z) := 2* x13. Definition P_id_A__U42 (x13:Z) (x14:Z) := 1* x14. Definition P_id_A__U72 (x13:Z) := 0. Definition P_id_A__U21 (x13:Z) (x14:Z) (x15:Z) := 1* x14 + 1* x15. Lemma P_id_A__U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U22 x14 x16 <= P_id_A__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNEPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNEPAL x14 <= P_id_A__ISNEPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U43 x14 <= P_id_A__U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U32 x14 <= P_id_A__U32 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U61 x14 x16 <= P_id_A__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U11 x14 x16 <= P_id_A__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U23 x14 <= P_id_A__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISPALLISTKIND_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPALLISTKIND x14 <= P_id_A__ISPALLISTKIND x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U52 x14 x16 <= P_id_A__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A____ x14 x16 <= P_id_A____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U41 x14 x16 x18 <= P_id_A__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U71 x14 x16 <= P_id_A__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNELIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNELIST x14 <= P_id_A__ISNELIST x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISLIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISLIST x14 <= P_id_A__ISLIST x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__AND x14 x16 <= P_id_A__AND x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U51 x14 x16 x18 <= P_id_A__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__ISQID_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISQID x14 <= P_id_A__ISQID x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U62 x14 <= P_id_A__U62 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U12 x14 <= P_id_A__U12 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U31 x14 x16 <= P_id_A__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISPAL x14 <= P_id_A__ISPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U53 x14 <= P_id_A__U53 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_MARK x14 <= P_id_MARK x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U42 x14 x16 <= P_id_A__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U72 x14 <= P_id_A__U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U21 x14 x16 x18 <= P_id_A__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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. Definition marked_measure := InterpZ.marked_measure 0 P_id_a____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal P_id_A__U22 P_id_A__ISNEPAL P_id_A__U43 P_id_A__U32 P_id_A__U61 P_id_A__U11 P_id_A__U23 P_id_A__ISPALLISTKIND P_id_A__U52 P_id_A____ P_id_A__U41 P_id_A__U71 P_id_A__ISNELIST P_id_A__ISLIST P_id_A__AND P_id_A__U51 P_id_A__ISQID P_id_A__U62 P_id_A__U12 P_id_A__U31 P_id_A__ISPAL P_id_A__U53 P_id_MARK P_id_A__U42 P_id_A__U72 P_id_A__U21. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_A__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_A__ISNEPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_A__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_A__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_A__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_A__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_A__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_A__ISPALLISTKIND (measure x13) | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_A__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_A____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15:: x14::x13::nil)) => P_id_A__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_A__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_A__ISNELIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_A__ISLIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_A__AND (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U51 (x15:: x14::x13::nil)) => P_id_A__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_A__ISQID (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_A__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_A__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_A__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_A__ISPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_A__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_MARK (measure x13) | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_A__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_A__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15:: x14::x13::nil)) => P_id_A__U21 (measure x15) (measure x14) (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__ISNEPAL_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U23_monotonic;assumption. intros ;apply P_id_A__ISPALLISTKIND_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A_____monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__U71_monotonic;assumption. intros ;apply P_id_A__ISNELIST_monotonic;assumption. intros ;apply P_id_A__ISLIST_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__ISQID_monotonic;assumption. intros ;apply P_id_A__U62_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__ISPAL_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__U72_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b) . Definition le a b := marked_measure a <= marked_measure b. Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c. Proof. unfold lt, le in *. intros a b c. apply (interp.le_lt_compat_right (interp.o_Z 0)). Qed. Lemma wf_lt : well_founded lt. Proof. unfold lt in *. apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. Qed. Lemma DP_R_xml_0_scc_23_large_large_large_large_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large_strict lt. Proof. unfold Relation_Definitions.inclusion, lt in *. intros a b H;destruct H; match goal with | |- (Zwf.Zwf 0) _ (marked_measure (algebra.Alg.Term ?f ?l)) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_lt_compat_right (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma DP_R_xml_0_scc_23_large_large_large_large_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large_large le. Proof. unfold Relation_Definitions.inclusion, le, Zwf.Zwf in *. intros a b H;destruct H; match goal with | |- _ <= marked_measure (algebra.Alg.Term ?f ?l) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_trans (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition wf_DP_R_xml_0_scc_23_large_large_large_large_large := WF_DP_R_xml_0_scc_23_large_large_large_large_large.wf. Lemma wf : well_founded WF_DP_R_xml_0_scc_23_large_large_large.DP_R_xml_0_scc_23_large_large_large_large . Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_23_large_large_large_large_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx'; apply DP_R_xml_0_scc_23_large_large_large_large_strict_in_lt; econstructor eassumption)|| ((apply IHx;[econstructor eassumption| intros y' Hlt;apply IHx';apply lt_le_compat with (1:=Hlt) ; apply DP_R_xml_0_scc_23_large_large_large_large_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_23_large_large_large_large_large. Qed. End WF_DP_R_xml_0_scc_23_large_large_large_large. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_a____ (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_a := 0. Definition P_id_a__U42 (x13:Z) (x14:Z) := 1* x13. Definition P_id_U42 (x13:Z) (x14:Z) := 1* x13. Definition P_id_a__U21 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_U21 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_a__U72 (x13:Z) := 0. Definition P_id_U72 (x13:Z) := 0. Definition P_id_a__U11 (x13:Z) (x14:Z) := 2* x13. Definition P_id_u := 0. Definition P_id_a__U53 (x13:Z) := 2* x13. Definition P_id_U53 (x13:Z) := 2* x13. Definition P_id_a__U31 (x13:Z) (x14:Z) := 2* x13. Definition P_id_U31 (x13:Z) (x14:Z) := 2* x13. Definition P_id_isPalListKind (x13:Z) := 0. Definition P_id_mark (x13:Z) := 1* x13. Definition P_id_i := 3. Definition P_id_a__U51 (x13:Z) (x14:Z) (x15:Z) := 1* x13. Definition P_id_U51 (x13:Z) (x14:Z) (x15:Z) := 1* x13. Definition P_id_a__isList (x13:Z) := 0. Definition P_id_isList (x13:Z) := 0. Definition P_id_a__and (x13:Z) (x14:Z) := 1* x13 + 2* x14. Definition P_id_a__U12 (x13:Z) := 1* x13. Definition P_id_U12 (x13:Z) := 1* x13. Definition P_id_a__U62 (x13:Z) := 1 + 1* x13. Definition P_id_U62 (x13:Z) := 1 + 1* x13. Definition P_id_a__isQid (x13:Z) := 0. Definition P_id_isQid (x13:Z) := 0. Definition P_id_isPal (x13:Z) := 0. Definition P_id___ (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_e := 0. Definition P_id_a__U43 (x13:Z) := 2* x13. Definition P_id_U43 (x13:Z) := 2* x13. Definition P_id_a__U22 (x13:Z) (x14:Z) := 2* x13. Definition P_id_U22 (x13:Z) (x14:Z) := 2* x13. Definition P_id_a__isNePal (x13:Z) := 1. Definition P_id_isNePal (x13:Z) := 1. Definition P_id_tt := 0. Definition P_id_U11 (x13:Z) (x14:Z) := 2* x13. Definition P_id_a__U61 (x13:Z) (x14:Z) := 1 + 1* x13. Definition P_id_U61 (x13:Z) (x14:Z) := 1 + 1* x13. Definition P_id_a__U32 (x13:Z) := 1* x13. Definition P_id_U32 (x13:Z) := 1* x13. Definition P_id_and (x13:Z) (x14:Z) := 1* x13 + 2* x14. Definition P_id_nil := 0. Definition P_id_o := 0. Definition P_id_a__U52 (x13:Z) (x14:Z) := 2* x13. Definition P_id_U52 (x13:Z) (x14:Z) := 2* x13. Definition P_id_a__U23 (x13:Z) := 2* x13. Definition P_id_U23 (x13:Z) := 2* x13. Definition P_id_a__isPalListKind (x13:Z) := 0. Definition P_id_a__isNeList (x13:Z) := 0. Definition P_id_isNeList (x13:Z) := 0. Definition P_id_a__U71 (x13:Z) (x14:Z) := 0. Definition P_id_U71 (x13:Z) (x14:Z) := 0. Definition P_id_a__U41 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_U41 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_a__isPal (x13:Z) := 0. Lemma P_id_a_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a____ x14 x16 <= P_id_a____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U42 x14 x16 <= P_id_a__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U42 x14 x16 <= P_id_U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U21 x14 x16 x18 <= P_id_a__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U21 x14 x16 x18 <= P_id_U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U72 x14 <= P_id_a__U72 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U72 x14 <= P_id_U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U11 x14 x16 <= P_id_a__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U53 x14 <= P_id_a__U53 x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U53 x14 <= P_id_U53 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U31 x14 x16 <= P_id_a__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U31 x14 x16 <= P_id_U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isPalListKind x14 <= P_id_isPalListKind x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_mark x14 <= P_id_mark x13. Proof. intros x14 x13. intros [H_1 H_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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U51 x14 x16 x18 <= P_id_a__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U51_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U51 x14 x16 x18 <= P_id_U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isList x14 <= P_id_a__isList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isList x14 <= P_id_isList x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__and x14 x16 <= P_id_a__and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U12_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U12 x14 <= P_id_a__U12 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U12 x14 <= P_id_U12 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U62 x14 <= P_id_a__U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U62 x14 <= P_id_U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isQid x14 <= P_id_a__isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isQid x14 <= P_id_isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isPal x14 <= P_id_isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id___ x14 x16 <= P_id___ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U43_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U43 x14 <= P_id_a__U43 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U43 x14 <= P_id_U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U22 x14 x16 <= P_id_a__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U22 x14 x16 <= P_id_U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNePal x14 <= P_id_a__isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isNePal x14 <= P_id_isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U11_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U11 x14 x16 <= P_id_U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U61 x14 x16 <= P_id_a__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U61 x14 x16 <= P_id_U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U32 x14 <= P_id_a__U32 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U32 x14 <= P_id_U32 x13. Proof. intros x14 x13. intros [H_1 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_and x14 x16 <= P_id_and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U52 x14 x16 <= P_id_a__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U52 x14 x16 <= P_id_U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U23 x14 <= P_id_a__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U23 x14 <= P_id_U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPalListKind x14 <= P_id_a__isPalListKind x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNeList x14 <= P_id_a__isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isNeList x14 <= P_id_isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U71 x14 x16 <= P_id_a__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U71 x14 x16 <= P_id_U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U41 x14 x16 x18 <= P_id_a__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U41_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U41 x14 x16 x18 <= P_id_U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isPal x14 <= P_id_a__isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a____ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U42 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U42 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U21 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U21 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U72 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_U72 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U11 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_u_bounded : 0 <= P_id_u . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U53 x13. 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 x13, (0 <= x13) ->0 <= P_id_U53 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U31 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U31 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPalListKind x13. 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 x13, (0 <= x13) ->0 <= P_id_mark x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_i_bounded : 0 <= P_id_i . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_bounded : forall x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U51 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U51 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isList x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__and x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U12 x13. 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 x13, (0 <= x13) ->0 <= P_id_U12 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id___ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_e_bounded : 0 <= P_id_e . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U43 x13. 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 x13, (0 <= x13) ->0 <= P_id_U43 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U22 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U22 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNePal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNePal x13. 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_U11_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U11 x14 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U61 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U61 x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U32 x13. 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 x13, (0 <= x13) ->0 <= P_id_U32 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_and x14 x13. 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_o_bounded : 0 <= P_id_o . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U52 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U52 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPalListKind x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U71 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U71 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U41 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U41 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPal x13. 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____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) => P_id_a____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_a__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U42 (x14::x13::nil)) => P_id_U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14:: x13::nil)) => P_id_a__U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U21 (x15::x14:: x13::nil)) => P_id_U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_a__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_U72 (x13::nil)) => P_id_U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_a__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_a__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_U53 (x13::nil)) => P_id_U53 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_a__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U31 (x14::x13::nil)) => P_id_U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil)) => P_id_isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_mark (measure x13) | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14:: x13::nil)) => P_id_a__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U51 (x15::x14:: x13::nil)) => P_id_U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_a__isList (measure x13) | (algebra.Alg.Term algebra.F.id_isList (x13::nil)) => P_id_isList (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_a__and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_a__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_U12 (x13::nil)) => P_id_U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_a__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_U62 (x13::nil)) => P_id_U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_a__isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isQid (x13::nil)) => P_id_isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isPal (x13::nil)) => P_id_isPal (measure x13) | (algebra.Alg.Term algebra.F.id___ (x14::x13::nil)) => P_id___ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_a__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_U43 (x13::nil)) => P_id_U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_a__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U22 (x14::x13::nil)) => P_id_U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_a__isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_isNePal (x13::nil)) => P_id_isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U11 (x14::x13::nil)) => P_id_U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_a__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U61 (x14::x13::nil)) => P_id_U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_a__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_U32 (x13::nil)) => P_id_U32 (measure x13) | (algebra.Alg.Term algebra.F.id_and (x14::x13::nil)) => P_id_and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_a__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U52 (x14::x13::nil)) => P_id_U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_a__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_U23 (x13::nil)) => P_id_U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_a__isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_a__isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_isNeList (x13::nil)) => P_id_isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_a__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U71 (x14::x13::nil)) => P_id_U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) => P_id_a__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U41 (x15::x14:: x13::nil)) => P_id_U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_a__isPal (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__U22 (x13:Z) (x14:Z) := 1. Definition P_id_A__ISNEPAL (x13:Z) := 2. Definition P_id_A__U43 (x13:Z) := 0. Definition P_id_A__U32 (x13:Z) := 0. Definition P_id_A__U61 (x13:Z) (x14:Z) := 0. Definition P_id_A__U11 (x13:Z) (x14:Z) := 1. Definition P_id_A__U23 (x13:Z) := 0. Definition P_id_A__ISPALLISTKIND (x13:Z) := 1. Definition P_id_A__U52 (x13:Z) (x14:Z) := 1. Definition P_id_A____ (x13:Z) (x14:Z) := 0. Definition P_id_A__U41 (x13:Z) (x14:Z) (x15:Z) := 1 + 2* x13. Definition P_id_A__U71 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNELIST (x13:Z) := 1. Definition P_id_A__ISLIST (x13:Z) := 1. Definition P_id_A__AND (x13:Z) (x14:Z) := 1 + 2* x14. Definition P_id_A__U51 (x13:Z) (x14:Z) (x15:Z) := 1. Definition P_id_A__ISQID (x13:Z) := 0. Definition P_id_A__U62 (x13:Z) := 0. Definition P_id_A__U12 (x13:Z) := 0. Definition P_id_A__U31 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISPAL (x13:Z) := 0. Definition P_id_A__U53 (x13:Z) := 0. Definition P_id_MARK (x13:Z) := 1 + 2* x13. Definition P_id_A__U42 (x13:Z) (x14:Z) := 1. Definition P_id_A__U72 (x13:Z) := 0. Definition P_id_A__U21 (x13:Z) (x14:Z) (x15:Z) := 1. Lemma P_id_A__U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U22 x14 x16 <= P_id_A__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNEPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNEPAL x14 <= P_id_A__ISNEPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U43 x14 <= P_id_A__U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U32 x14 <= P_id_A__U32 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U61 x14 x16 <= P_id_A__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U11 x14 x16 <= P_id_A__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U23 x14 <= P_id_A__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISPALLISTKIND_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPALLISTKIND x14 <= P_id_A__ISPALLISTKIND x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U52 x14 x16 <= P_id_A__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A____ x14 x16 <= P_id_A____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U41 x14 x16 x18 <= P_id_A__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U71 x14 x16 <= P_id_A__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNELIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNELIST x14 <= P_id_A__ISNELIST x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISLIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISLIST x14 <= P_id_A__ISLIST x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__AND x14 x16 <= P_id_A__AND x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U51 x14 x16 x18 <= P_id_A__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__ISQID_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISQID x14 <= P_id_A__ISQID x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U62 x14 <= P_id_A__U62 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U12 x14 <= P_id_A__U12 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U31 x14 x16 <= P_id_A__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISPAL x14 <= P_id_A__ISPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U53 x14 <= P_id_A__U53 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_MARK x14 <= P_id_MARK x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U42 x14 x16 <= P_id_A__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U72 x14 <= P_id_A__U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U21 x14 x16 x18 <= P_id_A__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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. Definition marked_measure := InterpZ.marked_measure 0 P_id_a____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal P_id_A__U22 P_id_A__ISNEPAL P_id_A__U43 P_id_A__U32 P_id_A__U61 P_id_A__U11 P_id_A__U23 P_id_A__ISPALLISTKIND P_id_A__U52 P_id_A____ P_id_A__U41 P_id_A__U71 P_id_A__ISNELIST P_id_A__ISLIST P_id_A__AND P_id_A__U51 P_id_A__ISQID P_id_A__U62 P_id_A__U12 P_id_A__U31 P_id_A__ISPAL P_id_A__U53 P_id_MARK P_id_A__U42 P_id_A__U72 P_id_A__U21. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_A__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_A__ISNEPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_A__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_A__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_A__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_A__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_A__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_A__ISPALLISTKIND (measure x13) | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_A__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_A____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15:: x14::x13::nil)) => P_id_A__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_A__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_A__ISNELIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_A__ISLIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_A__AND (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U51 (x15:: x14::x13::nil)) => P_id_A__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_A__ISQID (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_A__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_A__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_A__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_A__ISPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_A__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_MARK (measure x13) | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_A__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_A__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15:: x14::x13::nil)) => P_id_A__U21 (measure x15) (measure x14) (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__ISNEPAL_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U23_monotonic;assumption. intros ;apply P_id_A__ISPALLISTKIND_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A_____monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__U71_monotonic;assumption. intros ;apply P_id_A__ISNELIST_monotonic;assumption. intros ;apply P_id_A__ISLIST_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__ISQID_monotonic;assumption. intros ;apply P_id_A__U62_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__ISPAL_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__U72_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b). Definition le a b := marked_measure a <= marked_measure b. Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c. Proof. unfold lt, le in *. intros a b c. apply (interp.le_lt_compat_right (interp.o_Z 0)). Qed. Lemma wf_lt : well_founded lt. Proof. unfold lt in *. apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. Qed. Lemma DP_R_xml_0_scc_23_large_large_large_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_strict lt. Proof. unfold Relation_Definitions.inclusion, lt in *. intros a b H;destruct H; match goal with | |- (Zwf.Zwf 0) _ (marked_measure (algebra.Alg.Term ?f ?l)) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_lt_compat_right (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma DP_R_xml_0_scc_23_large_large_large_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large_large_large_large le. Proof. unfold Relation_Definitions.inclusion, le, Zwf.Zwf in *. intros a b H;destruct H; match goal with | |- _ <= marked_measure (algebra.Alg.Term ?f ?l) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_trans (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition wf_DP_R_xml_0_scc_23_large_large_large_large := WF_DP_R_xml_0_scc_23_large_large_large_large.wf. Lemma wf : well_founded WF_DP_R_xml_0_scc_23_large_large.DP_R_xml_0_scc_23_large_large_large . Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_23_large_large_large_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx';apply DP_R_xml_0_scc_23_large_large_large_strict_in_lt; econstructor eassumption)|| ((apply IHx;[econstructor eassumption| intros y' Hlt;apply IHx';apply lt_le_compat with (1:=Hlt) ; apply DP_R_xml_0_scc_23_large_large_large_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_23_large_large_large_large. Qed. End WF_DP_R_xml_0_scc_23_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____ (x13:Z) (x14:Z) := 2 + 2* x13 + 1* x14. Definition P_id_a := 0. Definition P_id_a__U42 (x13:Z) (x14:Z) := 2* x13. Definition P_id_U42 (x13:Z) (x14:Z) := 2* x13. Definition P_id_a__U21 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_U21 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_a__U72 (x13:Z) := 2 + 1* x13. Definition P_id_U72 (x13:Z) := 2 + 1* x13. Definition P_id_a__U11 (x13:Z) (x14:Z) := 1* x13. Definition P_id_u := 0. Definition P_id_a__U53 (x13:Z) := 2* x13. Definition P_id_U53 (x13:Z) := 2* x13. Definition P_id_a__U31 (x13:Z) (x14:Z) := 1* x13. Definition P_id_U31 (x13:Z) (x14:Z) := 1* x13. Definition P_id_isPalListKind (x13:Z) := 0. Definition P_id_mark (x13:Z) := 1* x13. Definition P_id_i := 1. Definition P_id_a__U51 (x13:Z) (x14:Z) (x15:Z) := 1* x13. Definition P_id_U51 (x13:Z) (x14:Z) (x15:Z) := 1* x13. Definition P_id_a__isList (x13:Z) := 0. Definition P_id_isList (x13:Z) := 0. Definition P_id_a__and (x13:Z) (x14:Z) := 2* x13 + 1* x14. Definition P_id_a__U12 (x13:Z) := 2* x13. Definition P_id_U12 (x13:Z) := 2* x13. Definition P_id_a__U62 (x13:Z) := 1* x13. Definition P_id_U62 (x13:Z) := 1* x13. Definition P_id_a__isQid (x13:Z) := 0. Definition P_id_isQid (x13:Z) := 0. Definition P_id_isPal (x13:Z) := 2 + 1* x13. Definition P_id___ (x13:Z) (x14:Z) := 2 + 2* x13 + 1* x14. Definition P_id_e := 0. Definition P_id_a__U43 (x13:Z) := 2* x13. Definition P_id_U43 (x13:Z) := 2* x13. Definition P_id_a__U22 (x13:Z) (x14:Z) := 2* x13. Definition P_id_U22 (x13:Z) (x14:Z) := 2* x13. Definition P_id_a__isNePal (x13:Z) := 1* x13. Definition P_id_isNePal (x13:Z) := 1* x13. Definition P_id_tt := 0. Definition P_id_U11 (x13:Z) (x14:Z) := 1* x13. Definition P_id_a__U61 (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_U61 (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_a__U32 (x13:Z) := 1* x13. Definition P_id_U32 (x13:Z) := 1* x13. Definition P_id_and (x13:Z) (x14:Z) := 2* x13 + 1* x14. Definition P_id_nil := 1. Definition P_id_o := 0. Definition P_id_a__U52 (x13:Z) (x14:Z) := 2* x13. Definition P_id_U52 (x13:Z) (x14:Z) := 2* x13. Definition P_id_a__U23 (x13:Z) := 2* x13. Definition P_id_U23 (x13:Z) := 2* x13. Definition P_id_a__isPalListKind (x13:Z) := 0. Definition P_id_a__isNeList (x13:Z) := 0. Definition P_id_isNeList (x13:Z) := 0. Definition P_id_a__U71 (x13:Z) (x14:Z) := 2 + 1* x13 + 1* x14. Definition P_id_U71 (x13:Z) (x14:Z) := 2 + 1* x13 + 1* x14. Definition P_id_a__U41 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_U41 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_a__isPal (x13:Z) := 2 + 1* x13. Lemma P_id_a_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a____ x14 x16 <= P_id_a____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U42 x14 x16 <= P_id_a__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U42 x14 x16 <= P_id_U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U21 x14 x16 x18 <= P_id_a__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U21 x14 x16 x18 <= P_id_U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U72 x14 <= P_id_a__U72 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U72 x14 <= P_id_U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U11 x14 x16 <= P_id_a__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U53 x14 <= P_id_a__U53 x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U53 x14 <= P_id_U53 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U31 x14 x16 <= P_id_a__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U31 x14 x16 <= P_id_U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isPalListKind x14 <= P_id_isPalListKind x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_mark x14 <= P_id_mark x13. Proof. intros x14 x13. intros [H_1 H_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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U51 x14 x16 x18 <= P_id_a__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U51_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U51 x14 x16 x18 <= P_id_U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isList x14 <= P_id_a__isList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isList x14 <= P_id_isList x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__and x14 x16 <= P_id_a__and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U12_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U12 x14 <= P_id_a__U12 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U12 x14 <= P_id_U12 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U62 x14 <= P_id_a__U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U62 x14 <= P_id_U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isQid x14 <= P_id_a__isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isQid x14 <= P_id_isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isPal x14 <= P_id_isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id___ x14 x16 <= P_id___ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U43_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U43 x14 <= P_id_a__U43 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U43 x14 <= P_id_U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U22 x14 x16 <= P_id_a__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U22 x14 x16 <= P_id_U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isNePal x14 <= P_id_a__isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isNePal x14 <= P_id_isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U11_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U11 x14 x16 <= P_id_U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U61 x14 x16 <= P_id_a__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U61 x14 x16 <= P_id_U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U32 x14 <= P_id_a__U32 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U32 x14 <= P_id_U32 x13. Proof. intros x14 x13. intros [H_1 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_and x14 x16 <= P_id_and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U52 x14 x16 <= P_id_a__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U52 x14 x16 <= P_id_U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U23 x14 <= P_id_a__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U23 x14 <= P_id_U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPalListKind x14 <= P_id_a__isPalListKind x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNeList x14 <= P_id_a__isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isNeList x14 <= P_id_isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U71 x14 x16 <= P_id_a__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U71 x14 x16 <= P_id_U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U41 x14 x16 x18 <= P_id_a__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U41_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U41 x14 x16 x18 <= P_id_U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isPal x14 <= P_id_a__isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a____ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U42 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U42 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U21 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U21 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U72 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_U72 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U11 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_u_bounded : 0 <= P_id_u . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U53 x13. 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 x13, (0 <= x13) ->0 <= P_id_U53 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U31 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U31 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPalListKind x13. 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 x13, (0 <= x13) ->0 <= P_id_mark x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_i_bounded : 0 <= P_id_i . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_bounded : forall x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U51 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U51 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isList x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__and x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U12 x13. 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 x13, (0 <= x13) ->0 <= P_id_U12 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id___ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_e_bounded : 0 <= P_id_e . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U43 x13. 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 x13, (0 <= x13) ->0 <= P_id_U43 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U22 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U22 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNePal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNePal x13. 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_U11_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U11 x14 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U61 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U61 x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U32 x13. 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 x13, (0 <= x13) ->0 <= P_id_U32 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_and x14 x13. 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_o_bounded : 0 <= P_id_o . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U52 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U52 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPalListKind x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U71 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U71 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U41 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U41 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPal x13. 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____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal . Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) => P_id_a____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) => P_id_a__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U42 (x14::x13::nil)) => P_id_U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14:: x13::nil)) => P_id_a__U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U21 (x15::x14:: x13::nil)) => P_id_U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_a__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_U72 (x13::nil)) => P_id_U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) => P_id_a__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_a__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_U53 (x13::nil)) => P_id_U53 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) => P_id_a__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U31 (x14::x13::nil)) => P_id_U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil)) => P_id_isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_mark (measure x13) | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14:: x13::nil)) => P_id_a__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U51 (x15::x14:: x13::nil)) => P_id_U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_a__isList (measure x13) | (algebra.Alg.Term algebra.F.id_isList (x13::nil)) => P_id_isList (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) => P_id_a__and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_a__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_U12 (x13::nil)) => P_id_U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_a__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_U62 (x13::nil)) => P_id_U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_a__isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isQid (x13::nil)) => P_id_isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isPal (x13::nil)) => P_id_isPal (measure x13) | (algebra.Alg.Term algebra.F.id___ (x14::x13::nil)) => P_id___ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_a__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_U43 (x13::nil)) => P_id_U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) => P_id_a__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U22 (x14::x13::nil)) => P_id_U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_a__isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_isNePal (x13::nil)) => P_id_isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U11 (x14::x13::nil)) => P_id_U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) => P_id_a__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U61 (x14::x13::nil)) => P_id_U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_a__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_U32 (x13::nil)) => P_id_U32 (measure x13) | (algebra.Alg.Term algebra.F.id_and (x14::x13::nil)) => P_id_and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o | (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) => P_id_a__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U52 (x14::x13::nil)) => P_id_U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_a__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_U23 (x13::nil)) => P_id_U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_a__isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_a__isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_isNeList (x13::nil)) => P_id_isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) => P_id_a__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U71 (x14::x13::nil)) => P_id_U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) => P_id_a__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U41 (x15::x14:: x13::nil)) => P_id_U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_a__isPal (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__U22 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNEPAL (x13:Z) := 2* x13. Definition P_id_A__U43 (x13:Z) := 0. Definition P_id_A__U32 (x13:Z) := 0. Definition P_id_A__U61 (x13:Z) (x14:Z) := 0. Definition P_id_A__U11 (x13:Z) (x14:Z) := 0. Definition P_id_A__U23 (x13:Z) := 0. Definition P_id_A__ISPALLISTKIND (x13:Z) := 0. Definition P_id_A__U52 (x13:Z) (x14:Z) := 0. Definition P_id_A____ (x13:Z) (x14:Z) := 0. Definition P_id_A__U41 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_A__U71 (x13:Z) (x14:Z) := 2 + 2* x14. Definition P_id_A__ISNELIST (x13:Z) := 0. Definition P_id_A__ISLIST (x13:Z) := 0. Definition P_id_A__AND (x13:Z) (x14:Z) := 2* x14. Definition P_id_A__U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_A__ISQID (x13:Z) := 0. Definition P_id_A__U62 (x13:Z) := 0. Definition P_id_A__U12 (x13:Z) := 0. Definition P_id_A__U31 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISPAL (x13:Z) := 3 + 2* x13. Definition P_id_A__U53 (x13:Z) := 0. Definition P_id_MARK (x13:Z) := 2* x13. Definition P_id_A__U42 (x13:Z) (x14:Z) := 0. Definition P_id_A__U72 (x13:Z) := 0. Definition P_id_A__U21 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Lemma P_id_A__U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U22 x14 x16 <= P_id_A__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNEPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISNEPAL x14 <= P_id_A__ISNEPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U43 x14 <= P_id_A__U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U32 x14 <= P_id_A__U32 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U61 x14 x16 <= P_id_A__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U11 x14 x16 <= P_id_A__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U23 x14 <= P_id_A__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISPALLISTKIND_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPALLISTKIND x14 <= P_id_A__ISPALLISTKIND x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U52 x14 x16 <= P_id_A__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A____ x14 x16 <= P_id_A____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U41 x14 x16 x18 <= P_id_A__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U71 x14 x16 <= P_id_A__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNELIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNELIST x14 <= P_id_A__ISNELIST x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISLIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISLIST x14 <= P_id_A__ISLIST x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__AND x14 x16 <= P_id_A__AND x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U51 x14 x16 x18 <= P_id_A__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__ISQID_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISQID x14 <= P_id_A__ISQID x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U62 x14 <= P_id_A__U62 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U12 x14 <= P_id_A__U12 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U31 x14 x16 <= P_id_A__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISPAL x14 <= P_id_A__ISPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U53 x14 <= P_id_A__U53 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_MARK x14 <= P_id_MARK x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U42 x14 x16 <= P_id_A__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U72 x14 <= P_id_A__U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U21 x14 x16 x18 <= P_id_A__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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. Definition marked_measure := InterpZ.marked_measure 0 P_id_a____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal P_id_A__U22 P_id_A__ISNEPAL P_id_A__U43 P_id_A__U32 P_id_A__U61 P_id_A__U11 P_id_A__U23 P_id_A__ISPALLISTKIND P_id_A__U52 P_id_A____ P_id_A__U41 P_id_A__U71 P_id_A__ISNELIST P_id_A__ISLIST P_id_A__AND P_id_A__U51 P_id_A__ISQID P_id_A__U62 P_id_A__U12 P_id_A__U31 P_id_A__ISPAL P_id_A__U53 P_id_MARK P_id_A__U42 P_id_A__U72 P_id_A__U21. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_A__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_A__ISNEPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_A__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_A__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_A__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_A__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_A__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_A__ISPALLISTKIND (measure x13) | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_A__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_A____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15:: x14::x13::nil)) => P_id_A__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_A__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_A__ISNELIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_A__ISLIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_A__AND (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U51 (x15:: x14::x13::nil)) => P_id_A__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_A__ISQID (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_A__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_A__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_A__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_A__ISPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_A__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_MARK (measure x13) | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_A__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_A__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15:: x14::x13::nil)) => P_id_A__U21 (measure x15) (measure x14) (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__ISNEPAL_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U23_monotonic;assumption. intros ;apply P_id_A__ISPALLISTKIND_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A_____monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__U71_monotonic;assumption. intros ;apply P_id_A__ISNELIST_monotonic;assumption. intros ;apply P_id_A__ISLIST_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__ISQID_monotonic;assumption. intros ;apply P_id_A__U62_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__ISPAL_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__U72_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b). Definition le a b := marked_measure a <= marked_measure b. Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c. Proof. unfold lt, le in *. intros a b c. apply (interp.le_lt_compat_right (interp.o_Z 0)). Qed. Lemma wf_lt : well_founded lt. Proof. unfold lt in *. apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. Qed. Lemma DP_R_xml_0_scc_23_large_large_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_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_23_large_large_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_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_23_large_large_large := WF_DP_R_xml_0_scc_23_large_large_large.wf. Lemma wf : well_founded WF_DP_R_xml_0_scc_23_large.DP_R_xml_0_scc_23_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_23_large_large_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx';apply DP_R_xml_0_scc_23_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_23_large_large_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_23_large_large_large. Qed. End WF_DP_R_xml_0_scc_23_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____ (x13:Z) (x14:Z) := 2 + 2* x13 + 1* x14. Definition P_id_a := 0. Definition P_id_a__U42 (x13:Z) (x14:Z) := 1* x13. Definition P_id_U42 (x13:Z) (x14:Z) := 1* x13. Definition P_id_a__U21 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_U21 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_a__U72 (x13:Z) := 2* x13. Definition P_id_U72 (x13:Z) := 2* x13. Definition P_id_a__U11 (x13:Z) (x14:Z) := 1* x13. Definition P_id_u := 1. Definition P_id_a__U53 (x13:Z) := 2* x13. Definition P_id_U53 (x13:Z) := 2* x13. Definition P_id_a__U31 (x13:Z) (x14:Z) := 1* x13. Definition P_id_U31 (x13:Z) (x14:Z) := 1* x13. Definition P_id_isPalListKind (x13:Z) := 0. Definition P_id_mark (x13:Z) := 1* x13. Definition P_id_i := 0. Definition P_id_a__U51 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_U51 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_a__isList (x13:Z) := 0. Definition P_id_isList (x13:Z) := 0. Definition P_id_a__and (x13:Z) (x14:Z) := 2* x13 + 2* x14. Definition P_id_a__U12 (x13:Z) := 2* x13. Definition P_id_U12 (x13:Z) := 2* x13. Definition P_id_a__U62 (x13:Z) := 1* x13. Definition P_id_U62 (x13:Z) := 1* x13. Definition P_id_a__isQid (x13:Z) := 0. Definition P_id_isQid (x13:Z) := 0. Definition P_id_isPal (x13:Z) := 0. Definition P_id___ (x13:Z) (x14:Z) := 2 + 2* x13 + 1* x14. Definition P_id_e := 0. Definition P_id_a__U43 (x13:Z) := 1* x13. Definition P_id_U43 (x13:Z) := 1* x13. Definition P_id_a__U22 (x13:Z) (x14:Z) := 1* x13. Definition P_id_U22 (x13:Z) (x14:Z) := 1* x13. Definition P_id_a__isNePal (x13:Z) := 0. Definition P_id_isNePal (x13:Z) := 0. Definition P_id_tt := 0. Definition P_id_U11 (x13:Z) (x14:Z) := 1* x13. Definition P_id_a__U61 (x13:Z) (x14:Z) := 2* x13. Definition P_id_U61 (x13:Z) (x14:Z) := 2* x13. Definition P_id_a__U32 (x13:Z) := 2* x13. Definition P_id_U32 (x13:Z) := 2* x13. Definition P_id_and (x13:Z) (x14:Z) := 2* x13 + 2* x14. Definition P_id_nil := 0. Definition P_id_o := 0. Definition P_id_a__U52 (x13:Z) (x14:Z) := 1* x13. Definition P_id_U52 (x13:Z) (x14:Z) := 1* x13. Definition P_id_a__U23 (x13:Z) := 2* x13. Definition P_id_U23 (x13:Z) := 2* x13. Definition P_id_a__isPalListKind (x13:Z) := 0. Definition P_id_a__isNeList (x13:Z) := 0. Definition P_id_isNeList (x13:Z) := 0. Definition P_id_a__U71 (x13:Z) (x14:Z) := 1* x13. Definition P_id_U71 (x13:Z) (x14:Z) := 1* x13. Definition P_id_a__U41 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_U41 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_a__isPal (x13:Z) := 0. Lemma P_id_a_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a____ x14 x16 <= P_id_a____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a__U42 x14 x16 <= P_id_a__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U42 x14 x16 <= P_id_U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U21 x14 x16 x18 <= P_id_a__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U21 x14 x16 x18 <= P_id_U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U72 x14 <= P_id_a__U72 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U72 x14 <= P_id_U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a__U11 x14 x16 <= P_id_a__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U53 x14 <= P_id_a__U53 x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U53 x14 <= P_id_U53 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a__U31 x14 x16 <= P_id_a__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U31 x14 x16 <= P_id_U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isPalListKind x14 <= P_id_isPalListKind x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_mark x14 <= P_id_mark x13. Proof. intros x14 x13. intros [H_1 H_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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U51 x14 x16 x18 <= P_id_a__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U51_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U51 x14 x16 x18 <= P_id_U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isList x14 <= P_id_a__isList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isList x14 <= P_id_isList x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a__and x14 x16 <= P_id_a__and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U12_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U12 x14 <= P_id_a__U12 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U12 x14 <= P_id_U12 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U62 x14 <= P_id_a__U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U62 x14 <= P_id_U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isQid x14 <= P_id_a__isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isQid x14 <= P_id_isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isPal x14 <= P_id_isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id___ x14 x16 <= P_id___ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U43_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U43 x14 <= P_id_a__U43 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U43 x14 <= P_id_U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a__U22 x14 x16 <= P_id_a__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U22 x14 x16 <= P_id_U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isNePal x14 <= P_id_a__isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isNePal x14 <= P_id_isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U11_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U11 x14 x16 <= P_id_U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a__U61 x14 x16 <= P_id_a__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U61 x14 x16 <= P_id_U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U32 x14 <= P_id_a__U32 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U32 x14 <= P_id_U32 x13. Proof. intros x14 x13. intros [H_1 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_and x14 x16 <= P_id_and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a__U52 x14 x16 <= P_id_a__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U52 x14 x16 <= P_id_U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U23 x14 <= P_id_a__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U23 x14 <= P_id_U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPalListKind x14 <= P_id_a__isPalListKind x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isNeList x14 <= P_id_a__isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isNeList x14 <= P_id_isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a__U71 x14 x16 <= P_id_a__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U71 x14 x16 <= P_id_U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U41 x14 x16 x18 <= P_id_a__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U41_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U41 x14 x16 x18 <= P_id_U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isPal x14 <= P_id_a__isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a____ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U42 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U42 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U21 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U21 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U72 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_U72 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U11 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_u_bounded : 0 <= P_id_u . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U53 x13. 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 x13, (0 <= x13) ->0 <= P_id_U53 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U31 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U31 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPalListKind x13. 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 x13, (0 <= x13) ->0 <= P_id_mark x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_i_bounded : 0 <= P_id_i . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_bounded : forall x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U51 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U51 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isList x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__and x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U12 x13. 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 x13, (0 <= x13) ->0 <= P_id_U12 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id___ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_e_bounded : 0 <= P_id_e . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U43 x13. 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 x13, (0 <= x13) ->0 <= P_id_U43 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U22 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U22 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNePal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNePal x13. 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_U11_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U11 x14 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U61 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U61 x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U32 x13. 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 x13, (0 <= x13) ->0 <= P_id_U32 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_and x14 x13. 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_o_bounded : 0 <= P_id_o . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U52 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U52 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPalListKind x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U71 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U71 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U41 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U41 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPal x13. 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____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) => P_id_a____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) => P_id_a__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U42 (x14::x13::nil)) => P_id_U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14:: x13::nil)) => P_id_a__U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U21 (x15::x14:: x13::nil)) => P_id_U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_a__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_U72 (x13::nil)) => P_id_U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) => P_id_a__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_a__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_U53 (x13::nil)) => P_id_U53 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) => P_id_a__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U31 (x14::x13::nil)) => P_id_U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil)) => P_id_isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_mark (measure x13) | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14:: x13::nil)) => P_id_a__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U51 (x15::x14:: x13::nil)) => P_id_U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_a__isList (measure x13) | (algebra.Alg.Term algebra.F.id_isList (x13::nil)) => P_id_isList (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) => P_id_a__and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_a__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_U12 (x13::nil)) => P_id_U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_a__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_U62 (x13::nil)) => P_id_U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_a__isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isQid (x13::nil)) => P_id_isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isPal (x13::nil)) => P_id_isPal (measure x13) | (algebra.Alg.Term algebra.F.id___ (x14::x13::nil)) => P_id___ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_a__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_U43 (x13::nil)) => P_id_U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) => P_id_a__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U22 (x14::x13::nil)) => P_id_U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_a__isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_isNePal (x13::nil)) => P_id_isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U11 (x14::x13::nil)) => P_id_U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) => P_id_a__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U61 (x14::x13::nil)) => P_id_U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_a__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_U32 (x13::nil)) => P_id_U32 (measure x13) | (algebra.Alg.Term algebra.F.id_and (x14::x13::nil)) => P_id_and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o | (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) => P_id_a__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U52 (x14::x13::nil)) => P_id_U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_a__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_U23 (x13::nil)) => P_id_U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_a__isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_a__isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_isNeList (x13::nil)) => P_id_isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) => P_id_a__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U71 (x14::x13::nil)) => P_id_U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) => P_id_a__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U41 (x15::x14:: x13::nil)) => P_id_U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_a__isPal (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__U22 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNEPAL (x13:Z) := 0. Definition P_id_A__U43 (x13:Z) := 0. Definition P_id_A__U32 (x13:Z) := 0. Definition P_id_A__U61 (x13:Z) (x14:Z) := 0. Definition P_id_A__U11 (x13:Z) (x14:Z) := 0. Definition P_id_A__U23 (x13:Z) := 0. Definition P_id_A__ISPALLISTKIND (x13:Z) := 0. Definition P_id_A__U52 (x13:Z) (x14:Z) := 1* x13. Definition P_id_A____ (x13:Z) (x14:Z) := 2* x13 + 1* x14. Definition P_id_A__U41 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_A__U71 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISNELIST (x13:Z) := 0. Definition P_id_A__ISLIST (x13:Z) := 0. Definition P_id_A__AND (x13:Z) (x14:Z) := 2* x14. Definition P_id_A__U51 (x13:Z) (x14:Z) (x15:Z) := 0. Definition P_id_A__ISQID (x13:Z) := 0. Definition P_id_A__U62 (x13:Z) := 0. Definition P_id_A__U12 (x13:Z) := 0. Definition P_id_A__U31 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISPAL (x13:Z) := 0. Definition P_id_A__U53 (x13:Z) := 0. Definition P_id_MARK (x13:Z) := 1* x13. Definition P_id_A__U42 (x13:Z) (x14:Z) := 0. Definition P_id_A__U72 (x13:Z) := 0. Definition P_id_A__U21 (x13:Z) (x14:Z) (x15:Z) := 0. Lemma P_id_A__U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A__U22 x14 x16 <= P_id_A__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNEPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISNEPAL x14 <= P_id_A__ISNEPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U43 x14 <= P_id_A__U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U32 x14 <= P_id_A__U32 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A__U61 x14 x16 <= P_id_A__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A__U11 x14 x16 <= P_id_A__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U23 x14 <= P_id_A__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISPALLISTKIND_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPALLISTKIND x14 <= P_id_A__ISPALLISTKIND x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A__U52 x14 x16 <= P_id_A__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A____ x14 x16 <= P_id_A____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U41 x14 x16 x18 <= P_id_A__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A__U71 x14 x16 <= P_id_A__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNELIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISNELIST x14 <= P_id_A__ISNELIST x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISLIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISLIST x14 <= P_id_A__ISLIST x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A__AND x14 x16 <= P_id_A__AND x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U51 x14 x16 x18 <= P_id_A__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__ISQID_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISQID x14 <= P_id_A__ISQID x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U62 x14 <= P_id_A__U62 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U12 x14 <= P_id_A__U12 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A__U31 x14 x16 <= P_id_A__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISPAL x14 <= P_id_A__ISPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U53 x14 <= P_id_A__U53 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_MARK x14 <= P_id_MARK x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A__U42 x14 x16 <= P_id_A__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U72 x14 <= P_id_A__U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U21 x14 x16 x18 <= P_id_A__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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. Definition marked_measure := InterpZ.marked_measure 0 P_id_a____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal P_id_A__U22 P_id_A__ISNEPAL P_id_A__U43 P_id_A__U32 P_id_A__U61 P_id_A__U11 P_id_A__U23 P_id_A__ISPALLISTKIND P_id_A__U52 P_id_A____ P_id_A__U41 P_id_A__U71 P_id_A__ISNELIST P_id_A__ISLIST P_id_A__AND P_id_A__U51 P_id_A__ISQID P_id_A__U62 P_id_A__U12 P_id_A__U31 P_id_A__ISPAL P_id_A__U53 P_id_MARK P_id_A__U42 P_id_A__U72 P_id_A__U21. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_A__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_A__ISNEPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_A__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_A__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_A__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_A__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_A__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_A__ISPALLISTKIND (measure x13) | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_A__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_A____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15:: x14::x13::nil)) => P_id_A__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_A__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_A__ISNELIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_A__ISLIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_A__AND (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U51 (x15:: x14::x13::nil)) => P_id_A__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_A__ISQID (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_A__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_A__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_A__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_A__ISPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_A__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_MARK (measure x13) | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_A__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_A__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15:: x14::x13::nil)) => P_id_A__U21 (measure x15) (measure x14) (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__ISNEPAL_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U23_monotonic;assumption. intros ;apply P_id_A__ISPALLISTKIND_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A_____monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__U71_monotonic;assumption. intros ;apply P_id_A__ISNELIST_monotonic;assumption. intros ;apply P_id_A__ISLIST_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__ISQID_monotonic;assumption. intros ;apply P_id_A__U62_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__ISPAL_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__U72_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b). Definition le a b := marked_measure a <= marked_measure b. Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c. Proof. unfold lt, le in *. intros a b c. apply (interp.le_lt_compat_right (interp.o_Z 0)). Qed. Lemma wf_lt : well_founded lt. Proof. unfold lt in *. apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. Qed. Lemma DP_R_xml_0_scc_23_large_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_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_23_large_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_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_23_large_large := WF_DP_R_xml_0_scc_23_large_large.wf. Lemma wf : well_founded WF_DP_R_xml_0_scc_23.DP_R_xml_0_scc_23_large. Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_23_large_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx';apply DP_R_xml_0_scc_23_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_23_large_large_in_le;econstructor eassumption])). apply wf_DP_R_xml_0_scc_23_large_large. Qed. End WF_DP_R_xml_0_scc_23_large. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_a____ (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_a := 0. Definition P_id_a__U42 (x13:Z) (x14:Z) := 2* x13. Definition P_id_U42 (x13:Z) (x14:Z) := 2* x13. Definition P_id_a__U21 (x13:Z) (x14:Z) (x15:Z) := 1* x13. Definition P_id_U21 (x13:Z) (x14:Z) (x15:Z) := 1* x13. Definition P_id_a__U72 (x13:Z) := 2* x13. Definition P_id_U72 (x13:Z) := 2* x13. Definition P_id_a__U11 (x13:Z) (x14:Z) := 2* x13. Definition P_id_u := 0. Definition P_id_a__U53 (x13:Z) := 1* x13. Definition P_id_U53 (x13:Z) := 1* x13. Definition P_id_a__U31 (x13:Z) (x14:Z) := 1* x13. Definition P_id_U31 (x13:Z) (x14:Z) := 1* x13. Definition P_id_isPalListKind (x13:Z) := 0. Definition P_id_mark (x13:Z) := 1* x13. Definition P_id_i := 3. Definition P_id_a__U51 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_U51 (x13:Z) (x14:Z) (x15:Z) := 2* x13. Definition P_id_a__isList (x13:Z) := 0. Definition P_id_isList (x13:Z) := 0. Definition P_id_a__and (x13:Z) (x14:Z) := 2* x13 + 2* x14. Definition P_id_a__U12 (x13:Z) := 2* x13. Definition P_id_U12 (x13:Z) := 2* x13. Definition P_id_a__U62 (x13:Z) := 1* x13. Definition P_id_U62 (x13:Z) := 1* x13. Definition P_id_a__isQid (x13:Z) := 0. Definition P_id_isQid (x13:Z) := 0. Definition P_id_isPal (x13:Z) := 0. Definition P_id___ (x13:Z) (x14:Z) := 1* x13 + 1* x14. Definition P_id_e := 0. Definition P_id_a__U43 (x13:Z) := 2* x13. Definition P_id_U43 (x13:Z) := 2* x13. Definition P_id_a__U22 (x13:Z) (x14:Z) := 1* x13. Definition P_id_U22 (x13:Z) (x14:Z) := 1* x13. Definition P_id_a__isNePal (x13:Z) := 0. Definition P_id_isNePal (x13:Z) := 0. Definition P_id_tt := 0. Definition P_id_U11 (x13:Z) (x14:Z) := 2* x13. Definition P_id_a__U61 (x13:Z) (x14:Z) := 1* x13. Definition P_id_U61 (x13:Z) (x14:Z) := 1* x13. Definition P_id_a__U32 (x13:Z) := 1* x13. Definition P_id_U32 (x13:Z) := 1* x13. Definition P_id_and (x13:Z) (x14:Z) := 2* x13 + 2* x14. Definition P_id_nil := 2. Definition P_id_o := 1. Definition P_id_a__U52 (x13:Z) (x14:Z) := 1* x13. Definition P_id_U52 (x13:Z) (x14:Z) := 1* x13. Definition P_id_a__U23 (x13:Z) := 1* x13. Definition P_id_U23 (x13:Z) := 1* x13. Definition P_id_a__isPalListKind (x13:Z) := 0. Definition P_id_a__isNeList (x13:Z) := 0. Definition P_id_isNeList (x13:Z) := 0. Definition P_id_a__U71 (x13:Z) (x14:Z) := 2* x13. Definition P_id_U71 (x13:Z) (x14:Z) := 2* x13. Definition P_id_a__U41 (x13:Z) (x14:Z) (x15:Z) := 1* x13. Definition P_id_U41 (x13:Z) (x14:Z) (x15:Z) := 1* x13. Definition P_id_a__isPal (x13:Z) := 0. Lemma P_id_a_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a____ x14 x16 <= P_id_a____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a__U42 x14 x16 <= P_id_a__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U42_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U42 x14 x16 <= P_id_U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U21 x14 x16 x18 <= P_id_a__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U21_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U21 x14 x16 x18 <= P_id_U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U72 x14 <= P_id_a__U72 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U72 x14 <= P_id_U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a__U11 x14 x16 <= P_id_a__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U53 x14 <= P_id_a__U53 x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U53 x14 <= P_id_U53 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a__U31 x14 x16 <= P_id_a__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U31 x14 x16 <= P_id_U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_isPalListKind x14 <= P_id_isPalListKind x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_mark x14 <= P_id_mark x13. Proof. intros x14 x13. intros [H_1 H_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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U51 x14 x16 x18 <= P_id_a__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U51_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U51 x14 x16 x18 <= P_id_U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isList x14 <= P_id_a__isList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isList x14 <= P_id_isList x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a__and x14 x16 <= P_id_a__and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U12_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U12 x14 <= P_id_a__U12 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U12 x14 <= P_id_U12 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U62 x14 <= P_id_a__U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U62 x14 <= P_id_U62 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isQid x14 <= P_id_a__isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isQid x14 <= P_id_isQid x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isPal x14 <= P_id_isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id___ x14 x16 <= P_id___ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U43_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U43 x14 <= P_id_a__U43 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U43 x14 <= P_id_U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a__U22 x14 x16 <= P_id_a__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U22 x14 x16 <= P_id_U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isNePal x14 <= P_id_a__isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isNePal x14 <= P_id_isNePal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U11_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U11 x14 x16 <= P_id_U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a__U61 x14 x16 <= P_id_a__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U61 x14 x16 <= P_id_U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U32 x14 <= P_id_a__U32 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U32 x14 <= P_id_U32 x13. Proof. intros x14 x13. intros [H_1 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_and x14 x16 <= P_id_and x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a__U52 x14 x16 <= P_id_a__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U52_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U52 x14 x16 <= P_id_U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__U23 x14 <= P_id_a__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_U23 x14 <= P_id_U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_a__isPalListKind x14 <= P_id_a__isPalListKind x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isNeList x14 <= P_id_a__isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_isNeList x14 <= P_id_isNeList x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_a__U71 x14 x16 <= P_id_a__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_U71 x14 x16 <= P_id_U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_a__U41 x14 x16 x18 <= P_id_a__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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_U41_monotonic : forall x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_U41 x14 x16 x18 <= P_id_U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__isPal_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_a__isPal x14 <= P_id_a__isPal x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a____ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U42_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U42 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U42 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U21 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U21 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U72 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U72_bounded : forall x13, (0 <= x13) ->0 <= P_id_U72 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U11 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_u_bounded : 0 <= P_id_u . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U53_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U53 x13. 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 x13, (0 <= x13) ->0 <= P_id_U53 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U31 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U31 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPalListKind x13. 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 x13, (0 <= x13) ->0 <= P_id_mark x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_i_bounded : 0 <= P_id_i . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U51_bounded : forall x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U51 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U51 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isList x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__and x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U12 x13. 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 x13, (0 <= x13) ->0 <= P_id_U12 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U62_bounded : forall x13, (0 <= x13) ->0 <= P_id_U62 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isQid_bounded : forall x13, (0 <= x13) ->0 <= P_id_isQid x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isPal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id____bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id___ x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_e_bounded : 0 <= P_id_e . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U43_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U43 x13. 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 x13, (0 <= x13) ->0 <= P_id_U43 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U22 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U22 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNePal x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNePal_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNePal x13. 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_U11_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U11 x14 x13. Proof. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U61 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U61 x14 x13. Proof. 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 x13, (0 <= x13) ->0 <= P_id_a__U32 x13. 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 x13, (0 <= x13) ->0 <= P_id_U32 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_and x14 x13. 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_o_bounded : 0 <= P_id_o . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U52_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U52 x14 x13. 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 x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U52 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U23_bounded : forall x13, (0 <= x13) ->0 <= P_id_U23 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPalListKind_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPalListKind x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_isNeList_bounded : forall x13, (0 <= x13) ->0 <= P_id_isNeList x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_a__U71 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_U71_bounded : forall x14 x13, (0 <= x13) ->(0 <= x14) ->0 <= P_id_U71 x14 x13. Proof. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_a__U41 x15 x14 x13. 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 x14 x13 x15, (0 <= x13) ->(0 <= x14) ->(0 <= x15) ->0 <= P_id_U41 x15 x14 x13. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a__isPal_bounded : forall x13, (0 <= x13) ->0 <= P_id_a__isPal x13. 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____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_a____ (x14::x13::nil)) => P_id_a____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_a__U42 (x14::x13::nil)) => P_id_a__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U42 (x14::x13::nil)) => P_id_U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14:: x13::nil)) => P_id_a__U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U21 (x15::x14::x13::nil)) => P_id_U21 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_a__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_U72 (x13::nil)) => P_id_U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14::x13::nil)) => P_id_a__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_u nil) => P_id_u | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_a__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_U53 (x13::nil)) => P_id_U53 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14::x13::nil)) => P_id_a__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U31 (x14::x13::nil)) => P_id_U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_isPalListKind (x13::nil)) => P_id_isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_mark (measure x13) | (algebra.Alg.Term algebra.F.id_i nil) => P_id_i | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14:: x13::nil)) => P_id_a__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U51 (x15::x14::x13::nil)) => P_id_U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_a__isList (measure x13) | (algebra.Alg.Term algebra.F.id_isList (x13::nil)) => P_id_isList (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14::x13::nil)) => P_id_a__and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_a__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_U12 (x13::nil)) => P_id_U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_a__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_U62 (x13::nil)) => P_id_U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_a__isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isQid (x13::nil)) => P_id_isQid (measure x13) | (algebra.Alg.Term algebra.F.id_isPal (x13::nil)) => P_id_isPal (measure x13) | (algebra.Alg.Term algebra.F.id___ (x14::x13::nil)) => P_id___ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_e nil) => P_id_e | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_a__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_U43 (x13::nil)) => P_id_U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U22 (x14::x13::nil)) => P_id_a__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U22 (x14::x13::nil)) => P_id_U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_a__isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_isNePal (x13::nil)) => P_id_isNePal (measure x13) | (algebra.Alg.Term algebra.F.id_tt nil) => P_id_tt | (algebra.Alg.Term algebra.F.id_U11 (x14::x13::nil)) => P_id_U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14::x13::nil)) => P_id_a__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U61 (x14::x13::nil)) => P_id_U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_a__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_U32 (x13::nil)) => P_id_U32 (measure x13) | (algebra.Alg.Term algebra.F.id_and (x14::x13::nil)) => P_id_and (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_o nil) => P_id_o | (algebra.Alg.Term algebra.F.id_a__U52 (x14::x13::nil)) => P_id_a__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U52 (x14::x13::nil)) => P_id_U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_a__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_U23 (x13::nil)) => P_id_U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_a__isPalListKind (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_a__isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_isNeList (x13::nil)) => P_id_isNeList (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14::x13::nil)) => P_id_a__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U71 (x14::x13::nil)) => P_id_U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) => P_id_a__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_U41 (x15::x14::x13::nil)) => P_id_U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_a__isPal (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_A__U22 (x13:Z) (x14:Z) := 1. Definition P_id_A__ISNEPAL (x13:Z) := 1. Definition P_id_A__U43 (x13:Z) := 0. Definition P_id_A__U32 (x13:Z) := 0. Definition P_id_A__U61 (x13:Z) (x14:Z) := 0. Definition P_id_A__U11 (x13:Z) (x14:Z) := 1. Definition P_id_A__U23 (x13:Z) := 0. Definition P_id_A__ISPALLISTKIND (x13:Z) := 1. Definition P_id_A__U52 (x13:Z) (x14:Z) := 1. Definition P_id_A____ (x13:Z) (x14:Z) := 1 + 1* x13 + 1* x14. Definition P_id_A__U41 (x13:Z) (x14:Z) (x15:Z) := 1. Definition P_id_A__U71 (x13:Z) (x14:Z) := 1. Definition P_id_A__ISNELIST (x13:Z) := 1. Definition P_id_A__ISLIST (x13:Z) := 1. Definition P_id_A__AND (x13:Z) (x14:Z) := 1 + 1* x14. Definition P_id_A__U51 (x13:Z) (x14:Z) (x15:Z) := 1 + 2* x13. Definition P_id_A__ISQID (x13:Z) := 0. Definition P_id_A__U62 (x13:Z) := 0. Definition P_id_A__U12 (x13:Z) := 0. Definition P_id_A__U31 (x13:Z) (x14:Z) := 0. Definition P_id_A__ISPAL (x13:Z) := 1. Definition P_id_A__U53 (x13:Z) := 0. Definition P_id_MARK (x13:Z) := 1 + 1* x13. Definition P_id_A__U42 (x13:Z) (x14:Z) := 1. Definition P_id_A__U72 (x13:Z) := 0. Definition P_id_A__U21 (x13:Z) (x14:Z) (x15:Z) := 1. Lemma P_id_A__U22_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A__U22 x14 x16 <= P_id_A__U22 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNEPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISNEPAL x14 <= P_id_A__ISNEPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U43 x14 <= P_id_A__U43 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U32 x14 <= P_id_A__U32 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A__U61 x14 x16 <= P_id_A__U61 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A__U11 x14 x16 <= P_id_A__U11 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U23_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U23 x14 <= P_id_A__U23 x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISPALLISTKIND_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) -> P_id_A__ISPALLISTKIND x14 <= P_id_A__ISPALLISTKIND x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A__U52 x14 x16 <= P_id_A__U52 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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_____monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A____ x14 x16 <= P_id_A____ x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U41 x14 x16 x18 <= P_id_A__U41 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__U71_monotonic : forall x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A__U71 x14 x16 <= P_id_A__U71 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISNELIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISNELIST x14 <= P_id_A__ISNELIST x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__ISLIST_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISLIST x14 <= P_id_A__ISLIST x13. Proof. intros x14 x13. intros [H_1 H_0]. 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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A__AND x14 x16 <= P_id_A__AND x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U51 x14 x16 x18 <= P_id_A__U51 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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__ISQID_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISQID x14 <= P_id_A__ISQID x13. Proof. intros x14 x13. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_A__U62_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U62 x14 <= P_id_A__U62 x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U12 x14 <= P_id_A__U12 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A__U31 x14 x16 <= P_id_A__U31 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__ISPAL_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__ISPAL x14 <= P_id_A__ISPAL x13. Proof. intros x14 x13. intros [H_1 H_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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U53 x14 <= P_id_A__U53 x13. Proof. intros x14 x13. intros [H_1 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 x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_MARK x14 <= P_id_MARK x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x14 x13 x15, (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) ->P_id_A__U42 x14 x16 <= P_id_A__U42 x13 x15. Proof. intros x16 x15 x14 x13. intros [H_1 H_0]. intros [H_3 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__U72_monotonic : forall x14 x13, (0 <= x14)/\ (x14 <= x13) ->P_id_A__U72 x14 <= P_id_A__U72 x13. Proof. intros x14 x13. intros [H_1 H_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 x16 x18 x14 x17 x13 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> (0 <= x14)/\ (x14 <= x13) -> P_id_A__U21 x14 x16 x18 <= P_id_A__U21 x13 x15 x17. Proof. intros x18 x17 x16 x15 x14 x13. 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. Definition marked_measure := InterpZ.marked_measure 0 P_id_a____ P_id_a P_id_a__U42 P_id_U42 P_id_a__U21 P_id_U21 P_id_a__U72 P_id_U72 P_id_a__U11 P_id_u P_id_a__U53 P_id_U53 P_id_a__U31 P_id_U31 P_id_isPalListKind P_id_mark P_id_i P_id_a__U51 P_id_U51 P_id_a__isList P_id_isList P_id_a__and P_id_a__U12 P_id_U12 P_id_a__U62 P_id_U62 P_id_a__isQid P_id_isQid P_id_isPal P_id___ P_id_e P_id_a__U43 P_id_U43 P_id_a__U22 P_id_U22 P_id_a__isNePal P_id_isNePal P_id_tt P_id_U11 P_id_a__U61 P_id_U61 P_id_a__U32 P_id_U32 P_id_and P_id_nil P_id_o P_id_a__U52 P_id_U52 P_id_a__U23 P_id_U23 P_id_a__isPalListKind P_id_a__isNeList P_id_isNeList P_id_a__U71 P_id_U71 P_id_a__U41 P_id_U41 P_id_a__isPal P_id_A__U22 P_id_A__ISNEPAL P_id_A__U43 P_id_A__U32 P_id_A__U61 P_id_A__U11 P_id_A__U23 P_id_A__ISPALLISTKIND P_id_A__U52 P_id_A____ P_id_A__U41 P_id_A__U71 P_id_A__ISNELIST P_id_A__ISLIST P_id_A__AND P_id_A__U51 P_id_A__ISQID P_id_A__U62 P_id_A__U12 P_id_A__U31 P_id_A__ISPAL P_id_A__U53 P_id_MARK P_id_A__U42 P_id_A__U72 P_id_A__U21 . Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_a__U22 (x14:: x13::nil)) => P_id_A__U22 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNePal (x13::nil)) => P_id_A__ISNEPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U43 (x13::nil)) => P_id_A__U43 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U32 (x13::nil)) => P_id_A__U32 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U61 (x14:: x13::nil)) => P_id_A__U61 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U11 (x14:: x13::nil)) => P_id_A__U11 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U23 (x13::nil)) => P_id_A__U23 (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPalListKind (x13::nil)) => P_id_A__ISPALLISTKIND (measure x13) | (algebra.Alg.Term algebra.F.id_a__U52 (x14:: x13::nil)) => P_id_A__U52 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a____ (x14:: x13::nil)) => P_id_A____ (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U41 (x15::x14:: x13::nil)) => P_id_A__U41 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U71 (x14:: x13::nil)) => P_id_A__U71 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isNeList (x13::nil)) => P_id_A__ISNELIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__isList (x13::nil)) => P_id_A__ISLIST (measure x13) | (algebra.Alg.Term algebra.F.id_a__and (x14:: x13::nil)) => P_id_A__AND (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U51 (x15::x14:: x13::nil)) => P_id_A__U51 (measure x15) (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isQid (x13::nil)) => P_id_A__ISQID (measure x13) | (algebra.Alg.Term algebra.F.id_a__U62 (x13::nil)) => P_id_A__U62 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U12 (x13::nil)) => P_id_A__U12 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U31 (x14:: x13::nil)) => P_id_A__U31 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__isPal (x13::nil)) => P_id_A__ISPAL (measure x13) | (algebra.Alg.Term algebra.F.id_a__U53 (x13::nil)) => P_id_A__U53 (measure x13) | (algebra.Alg.Term algebra.F.id_mark (x13::nil)) => P_id_MARK (measure x13) | (algebra.Alg.Term algebra.F.id_a__U42 (x14:: x13::nil)) => P_id_A__U42 (measure x14) (measure x13) | (algebra.Alg.Term algebra.F.id_a__U72 (x13::nil)) => P_id_A__U72 (measure x13) | (algebra.Alg.Term algebra.F.id_a__U21 (x15::x14:: x13::nil)) => P_id_A__U21 (measure x15) (measure x14) (measure x13) | _ => 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_a_____monotonic;assumption. intros ;apply P_id_a__U42_monotonic;assumption. intros ;apply P_id_U42_monotonic;assumption. intros ;apply P_id_a__U21_monotonic;assumption. intros ;apply P_id_U21_monotonic;assumption. intros ;apply P_id_a__U72_monotonic;assumption. intros ;apply P_id_U72_monotonic;assumption. intros ;apply P_id_a__U11_monotonic;assumption. intros ;apply P_id_a__U53_monotonic;assumption. intros ;apply P_id_U53_monotonic;assumption. intros ;apply P_id_a__U31_monotonic;assumption. intros ;apply P_id_U31_monotonic;assumption. intros ;apply P_id_isPalListKind_monotonic;assumption. intros ;apply P_id_mark_monotonic;assumption. intros ;apply P_id_a__U51_monotonic;assumption. intros ;apply P_id_U51_monotonic;assumption. intros ;apply P_id_a__isList_monotonic;assumption. intros ;apply P_id_isList_monotonic;assumption. intros ;apply P_id_a__and_monotonic;assumption. intros ;apply P_id_a__U12_monotonic;assumption. intros ;apply P_id_U12_monotonic;assumption. intros ;apply P_id_a__U62_monotonic;assumption. intros ;apply P_id_U62_monotonic;assumption. intros ;apply P_id_a__isQid_monotonic;assumption. intros ;apply P_id_isQid_monotonic;assumption. intros ;apply P_id_isPal_monotonic;assumption. intros ;apply P_id____monotonic;assumption. intros ;apply P_id_a__U43_monotonic;assumption. intros ;apply P_id_U43_monotonic;assumption. intros ;apply P_id_a__U22_monotonic;assumption. intros ;apply P_id_U22_monotonic;assumption. intros ;apply P_id_a__isNePal_monotonic;assumption. intros ;apply P_id_isNePal_monotonic;assumption. intros ;apply P_id_U11_monotonic;assumption. intros ;apply P_id_a__U61_monotonic;assumption. intros ;apply P_id_U61_monotonic;assumption. intros ;apply P_id_a__U32_monotonic;assumption. intros ;apply P_id_U32_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_a__U52_monotonic;assumption. intros ;apply P_id_U52_monotonic;assumption. intros ;apply P_id_a__U23_monotonic;assumption. intros ;apply P_id_U23_monotonic;assumption. intros ;apply P_id_a__isPalListKind_monotonic;assumption. intros ;apply P_id_a__isNeList_monotonic;assumption. intros ;apply P_id_isNeList_monotonic;assumption. intros ;apply P_id_a__U71_monotonic;assumption. intros ;apply P_id_U71_monotonic;assumption. intros ;apply P_id_a__U41_monotonic;assumption. intros ;apply P_id_U41_monotonic;assumption. intros ;apply P_id_a__isPal_monotonic;assumption. intros ;apply P_id_a_____bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_a__U42_bounded;assumption. intros ;apply P_id_U42_bounded;assumption. intros ;apply P_id_a__U21_bounded;assumption. intros ;apply P_id_U21_bounded;assumption. intros ;apply P_id_a__U72_bounded;assumption. intros ;apply P_id_U72_bounded;assumption. intros ;apply P_id_a__U11_bounded;assumption. intros ;apply P_id_u_bounded;assumption. intros ;apply P_id_a__U53_bounded;assumption. intros ;apply P_id_U53_bounded;assumption. intros ;apply P_id_a__U31_bounded;assumption. intros ;apply P_id_U31_bounded;assumption. intros ;apply P_id_isPalListKind_bounded;assumption. intros ;apply P_id_mark_bounded;assumption. intros ;apply P_id_i_bounded;assumption. intros ;apply P_id_a__U51_bounded;assumption. intros ;apply P_id_U51_bounded;assumption. intros ;apply P_id_a__isList_bounded;assumption. intros ;apply P_id_isList_bounded;assumption. intros ;apply P_id_a__and_bounded;assumption. intros ;apply P_id_a__U12_bounded;assumption. intros ;apply P_id_U12_bounded;assumption. intros ;apply P_id_a__U62_bounded;assumption. intros ;apply P_id_U62_bounded;assumption. intros ;apply P_id_a__isQid_bounded;assumption. intros ;apply P_id_isQid_bounded;assumption. intros ;apply P_id_isPal_bounded;assumption. intros ;apply P_id____bounded;assumption. intros ;apply P_id_e_bounded;assumption. intros ;apply P_id_a__U43_bounded;assumption. intros ;apply P_id_U43_bounded;assumption. intros ;apply P_id_a__U22_bounded;assumption. intros ;apply P_id_U22_bounded;assumption. intros ;apply P_id_a__isNePal_bounded;assumption. intros ;apply P_id_isNePal_bounded;assumption. intros ;apply P_id_tt_bounded;assumption. intros ;apply P_id_U11_bounded;assumption. intros ;apply P_id_a__U61_bounded;assumption. intros ;apply P_id_U61_bounded;assumption. intros ;apply P_id_a__U32_bounded;assumption. intros ;apply P_id_U32_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_o_bounded;assumption. intros ;apply P_id_a__U52_bounded;assumption. intros ;apply P_id_U52_bounded;assumption. intros ;apply P_id_a__U23_bounded;assumption. intros ;apply P_id_U23_bounded;assumption. intros ;apply P_id_a__isPalListKind_bounded;assumption. intros ;apply P_id_a__isNeList_bounded;assumption. intros ;apply P_id_isNeList_bounded;assumption. intros ;apply P_id_a__U71_bounded;assumption. intros ;apply P_id_U71_bounded;assumption. intros ;apply P_id_a__U41_bounded;assumption. intros ;apply P_id_U41_bounded;assumption. intros ;apply P_id_a__isPal_bounded;assumption. apply rules_monotonic. intros ;apply P_id_A__U22_monotonic;assumption. intros ;apply P_id_A__ISNEPAL_monotonic;assumption. intros ;apply P_id_A__U43_monotonic;assumption. intros ;apply P_id_A__U32_monotonic;assumption. intros ;apply P_id_A__U61_monotonic;assumption. intros ;apply P_id_A__U11_monotonic;assumption. intros ;apply P_id_A__U23_monotonic;assumption. intros ;apply P_id_A__ISPALLISTKIND_monotonic;assumption. intros ;apply P_id_A__U52_monotonic;assumption. intros ;apply P_id_A_____monotonic;assumption. intros ;apply P_id_A__U41_monotonic;assumption. intros ;apply P_id_A__U71_monotonic;assumption. intros ;apply P_id_A__ISNELIST_monotonic;assumption. intros ;apply P_id_A__ISLIST_monotonic;assumption. intros ;apply P_id_A__AND_monotonic;assumption. intros ;apply P_id_A__U51_monotonic;assumption. intros ;apply P_id_A__ISQID_monotonic;assumption. intros ;apply P_id_A__U62_monotonic;assumption. intros ;apply P_id_A__U12_monotonic;assumption. intros ;apply P_id_A__U31_monotonic;assumption. intros ;apply P_id_A__ISPAL_monotonic;assumption. intros ;apply P_id_A__U53_monotonic;assumption. intros ;apply P_id_MARK_monotonic;assumption. intros ;apply P_id_A__U42_monotonic;assumption. intros ;apply P_id_A__U72_monotonic;assumption. intros ;apply P_id_A__U21_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b). Definition le a b := marked_measure a <= marked_measure b. Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c. Proof. unfold lt, le in *. intros a b c. apply (interp.le_lt_compat_right (interp.o_Z 0)). Qed. Lemma wf_lt : well_founded lt. Proof. unfold lt in *. apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. Qed. Lemma DP_R_xml_0_scc_23_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_strict lt. Proof. unfold Relation_Definitions.inclusion, lt in *. intros a b H;destruct H; match goal with | |- (Zwf.Zwf 0) _ (marked_measure (algebra.Alg.Term ?f ?l)) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_lt_compat_right (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma DP_R_xml_0_scc_23_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_23_large le. Proof. unfold Relation_Definitions.inclusion, le, Zwf.Zwf in *. intros a b H;destruct H; match goal with | |- _ <= marked_measure (algebra.Alg.Term ?f ?l) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_trans (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition wf_DP_R_xml_0_scc_23_large := WF_DP_R_xml_0_scc_23_large.wf. Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_23. Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_23_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx';apply DP_R_xml_0_scc_23_strict_in_lt; econstructor eassumption)|| ((apply IHx;[econstructor eassumption| intros y' Hlt;apply IHx';apply lt_le_compat with (1:=Hlt) ; apply DP_R_xml_0_scc_23_large_in_le;econstructor eassumption])). apply wf_DP_R_xml_0_scc_23_large. Qed. End WF_DP_R_xml_0_scc_23. Definition wf_DP_R_xml_0_scc_23 := WF_DP_R_xml_0_scc_23.wf. Lemma acc_DP_R_xml_0_scc_23 : forall x y, (DP_R_xml_0_scc_23 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_23). intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((eapply acc_DP_R_xml_0_non_scc_22; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_21; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_20; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_19; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_16; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_15; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_14; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_13; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_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_23. 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_22; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_21; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_20; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_19; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_18; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_17; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_16; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_15; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_14; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_13; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_12; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_11; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_10; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_9; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_8; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_7; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_6; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_5; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_4; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_22; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_21; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_20; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_19; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_18; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_17; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_16; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_15; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_14; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_13; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_11; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_10; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_9; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_8; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_7; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_6; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_5; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_4; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_3; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_2; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_1; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_0; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (fail)))))))))))))))))))))))))))))))))))))))))))))))). Qed. End WF_DP_R_xml_0. Definition wf_H := WF_DP_R_xml_0.wf. Lemma wf : well_founded (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules). Proof. apply ddp.dp_criterion. apply R_xml_0_deep_rew.R_xml_0_non_var. apply R_xml_0_deep_rew.R_xml_0_reg. intros ; apply (ddp.constructor_defined_dec _ _ R_xml_0_deep_rew.R_xml_0_rules_included). refine (Inclusion.wf_incl _ _ _ _ wf_H). intros x y H. destruct (R_xml_0_dp_step_spec H) as [f [l1 [l2 [H1 [H2 H3]]]]]. destruct (ddp.dp_list_complete _ _ R_xml_0_deep_rew.R_xml_0_rules_included _ _ H3) as [x' [y' [sigma [h1 [h2 h3]]]]]. clear H3. subst. vm_compute in h3|-. let e := type of h3 in (dp_concl_tac h2 h3 ltac:(fun _ => idtac) e). Qed. End WF_R_xml_0_deep_rew. (* *** Local Variables: *** *** coq-prog-name: "coqtop" *** *** coq-prog-args: ("-emacs-U" "-I" "$COCCINELLE/examples" "-I" "$COCCINELLE/term_algebra" "-I" "$COCCINELLE/term_orderings" "-I" "$COCCINELLE/basis" "-I" "$COCCINELLE/list_extensions" "-I" "$COCCINELLE/examples/cime_trace/") *** *** compile-command: "coqc -I $COCCINELLE/term_algebra -I $COCCINELLE/term_orderings -I $COCCINELLE/basis -I $COCCINELLE/list_extensions -I $COCCINELLE/examples/cime_trace/ -I $COCCINELLE/examples/ c_output/strat/tpdb-5.0___TRS___TRCSR___PALINDROME_complete_GM.trs/a3pat.v" *** *** End: *** *)