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_test *) | id_test : symb (* id_match_1 *) | id_match_1 : symb (* id_match_0 *) | id_match_0 : symb (* id_quick *) | id_quick : symb (* id_False *) | id_False : symb (* id_match_2 *) | id_match_2 : symb (* id_Cons *) | id_Cons : symb (* id_match_5 *) | id_match_5 : symb (* id_True *) | id_True : symb (* id_Pair *) | id_Pair : symb (* id_Nil *) | id_Nil : symb (* id_match_4 *) | id_match_4 : symb (* id_append *) | id_append : symb (* id_match_3 *) | id_match_3 : symb (* id_part *) | id_part : symb . Definition symb_eq_bool (f1 f2:symb) : bool := match f1,f2 with | id_test,id_test => true | id_match_1,id_match_1 => true | id_match_0,id_match_0 => true | id_quick,id_quick => true | id_False,id_False => true | id_match_2,id_match_2 => true | id_Cons,id_Cons => true | id_match_5,id_match_5 => true | id_True,id_True => true | id_Pair,id_Pair => true | id_Nil,id_Nil => true | id_match_4,id_match_4 => true | id_append,id_append => true | id_match_3,id_match_3 => true | id_part,id_part => 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_test,id_test => refl_equal _ | id_match_1,id_match_1 => refl_equal _ | id_match_0,id_match_0 => refl_equal _ | id_quick,id_quick => refl_equal _ | id_False,id_False => refl_equal _ | id_match_2,id_match_2 => refl_equal _ | id_Cons,id_Cons => refl_equal _ | id_match_5,id_match_5 => refl_equal _ | id_True,id_True => refl_equal _ | id_Pair,id_Pair => refl_equal _ | id_Nil,id_Nil => refl_equal _ | id_match_4,id_match_4 => refl_equal _ | id_append,id_append => refl_equal _ | id_match_3,id_match_3 => refl_equal _ | id_part,id_part => refl_equal _ | _,_ => _ end;intros abs;discriminate. Defined. Definition arity (f:symb) := match f with | id_test => term.Free 2 | id_match_1 => term.Free 3 | id_match_0 => term.Free 3 | id_quick => term.Free 1 | id_False => term.Free 0 | id_match_2 => term.Free 5 | id_Cons => term.Free 2 | id_match_5 => term.Free 4 | id_True => term.Free 0 | id_Pair => term.Free 2 | id_Nil => term.Free 0 | id_match_4 => term.Free 2 | id_append => term.Free 2 | id_match_3 => term.Free 7 | id_part => term.Free 2 end. Definition symb_order (f1 f2:symb) : bool := match f1,f2 with | id_test,id_test => true | id_test,id_match_1 => false | id_test,id_match_0 => false | id_test,id_quick => false | id_test,id_False => false | id_test,id_match_2 => false | id_test,id_Cons => false | id_test,id_match_5 => false | id_test,id_True => false | id_test,id_Pair => false | id_test,id_Nil => false | id_test,id_match_4 => false | id_test,id_append => false | id_test,id_match_3 => false | id_test,id_part => false | id_match_1,id_test => true | id_match_1,id_match_1 => true | id_match_1,id_match_0 => false | id_match_1,id_quick => false | id_match_1,id_False => false | id_match_1,id_match_2 => false | id_match_1,id_Cons => false | id_match_1,id_match_5 => false | id_match_1,id_True => false | id_match_1,id_Pair => false | id_match_1,id_Nil => false | id_match_1,id_match_4 => false | id_match_1,id_append => false | id_match_1,id_match_3 => false | id_match_1,id_part => false | id_match_0,id_test => true | id_match_0,id_match_1 => true | id_match_0,id_match_0 => true | id_match_0,id_quick => false | id_match_0,id_False => false | id_match_0,id_match_2 => false | id_match_0,id_Cons => false | id_match_0,id_match_5 => false | id_match_0,id_True => false | id_match_0,id_Pair => false | id_match_0,id_Nil => false | id_match_0,id_match_4 => false | id_match_0,id_append => false | id_match_0,id_match_3 => false | id_match_0,id_part => false | id_quick,id_test => true | id_quick,id_match_1 => true | id_quick,id_match_0 => true | id_quick,id_quick => true | id_quick,id_False => false | id_quick,id_match_2 => false | id_quick,id_Cons => false | id_quick,id_match_5 => false | id_quick,id_True => false | id_quick,id_Pair => false | id_quick,id_Nil => false | id_quick,id_match_4 => false | id_quick,id_append => false | id_quick,id_match_3 => false | id_quick,id_part => false | id_False,id_test => true | id_False,id_match_1 => true | id_False,id_match_0 => true | id_False,id_quick => true | id_False,id_False => true | id_False,id_match_2 => false | id_False,id_Cons => false | id_False,id_match_5 => false | id_False,id_True => false | id_False,id_Pair => false | id_False,id_Nil => false | id_False,id_match_4 => false | id_False,id_append => false | id_False,id_match_3 => false | id_False,id_part => false | id_match_2,id_test => true | id_match_2,id_match_1 => true | id_match_2,id_match_0 => true | id_match_2,id_quick => true | id_match_2,id_False => true | id_match_2,id_match_2 => true | id_match_2,id_Cons => false | id_match_2,id_match_5 => false | id_match_2,id_True => false | id_match_2,id_Pair => false | id_match_2,id_Nil => false | id_match_2,id_match_4 => false | id_match_2,id_append => false | id_match_2,id_match_3 => false | id_match_2,id_part => false | id_Cons,id_test => true | id_Cons,id_match_1 => true | id_Cons,id_match_0 => true | id_Cons,id_quick => true | id_Cons,id_False => true | id_Cons,id_match_2 => true | id_Cons,id_Cons => true | id_Cons,id_match_5 => false | id_Cons,id_True => false | id_Cons,id_Pair => false | id_Cons,id_Nil => false | id_Cons,id_match_4 => false | id_Cons,id_append => false | id_Cons,id_match_3 => false | id_Cons,id_part => false | id_match_5,id_test => true | id_match_5,id_match_1 => true | id_match_5,id_match_0 => true | id_match_5,id_quick => true | id_match_5,id_False => true | id_match_5,id_match_2 => true | id_match_5,id_Cons => true | id_match_5,id_match_5 => true | id_match_5,id_True => false | id_match_5,id_Pair => false | id_match_5,id_Nil => false | id_match_5,id_match_4 => false | id_match_5,id_append => false | id_match_5,id_match_3 => false | id_match_5,id_part => false | id_True,id_test => true | id_True,id_match_1 => true | id_True,id_match_0 => true | id_True,id_quick => true | id_True,id_False => true | id_True,id_match_2 => true | id_True,id_Cons => true | id_True,id_match_5 => true | id_True,id_True => true | id_True,id_Pair => false | id_True,id_Nil => false | id_True,id_match_4 => false | id_True,id_append => false | id_True,id_match_3 => false | id_True,id_part => false | id_Pair,id_test => true | id_Pair,id_match_1 => true | id_Pair,id_match_0 => true | id_Pair,id_quick => true | id_Pair,id_False => true | id_Pair,id_match_2 => true | id_Pair,id_Cons => true | id_Pair,id_match_5 => true | id_Pair,id_True => true | id_Pair,id_Pair => true | id_Pair,id_Nil => false | id_Pair,id_match_4 => false | id_Pair,id_append => false | id_Pair,id_match_3 => false | id_Pair,id_part => false | id_Nil,id_test => true | id_Nil,id_match_1 => true | id_Nil,id_match_0 => true | id_Nil,id_quick => true | id_Nil,id_False => true | id_Nil,id_match_2 => true | id_Nil,id_Cons => true | id_Nil,id_match_5 => true | id_Nil,id_True => true | id_Nil,id_Pair => true | id_Nil,id_Nil => true | id_Nil,id_match_4 => false | id_Nil,id_append => false | id_Nil,id_match_3 => false | id_Nil,id_part => false | id_match_4,id_test => true | id_match_4,id_match_1 => true | id_match_4,id_match_0 => true | id_match_4,id_quick => true | id_match_4,id_False => true | id_match_4,id_match_2 => true | id_match_4,id_Cons => true | id_match_4,id_match_5 => true | id_match_4,id_True => true | id_match_4,id_Pair => true | id_match_4,id_Nil => true | id_match_4,id_match_4 => true | id_match_4,id_append => false | id_match_4,id_match_3 => false | id_match_4,id_part => false | id_append,id_test => true | id_append,id_match_1 => true | id_append,id_match_0 => true | id_append,id_quick => true | id_append,id_False => true | id_append,id_match_2 => true | id_append,id_Cons => true | id_append,id_match_5 => true | id_append,id_True => true | id_append,id_Pair => true | id_append,id_Nil => true | id_append,id_match_4 => true | id_append,id_append => true | id_append,id_match_3 => false | id_append,id_part => false | id_match_3,id_test => true | id_match_3,id_match_1 => true | id_match_3,id_match_0 => true | id_match_3,id_quick => true | id_match_3,id_False => true | id_match_3,id_match_2 => true | id_match_3,id_Cons => true | id_match_3,id_match_5 => true | id_match_3,id_True => true | id_match_3,id_Pair => true | id_match_3,id_Nil => true | id_match_3,id_match_4 => true | id_match_3,id_append => true | id_match_3,id_match_3 => true | id_match_3,id_part => false | id_part,id_test => true | id_part,id_match_1 => true | id_part,id_match_0 => true | id_part,id_quick => true | id_part,id_False => true | id_part,id_match_2 => true | id_part,id_Cons => true | id_part,id_match_5 => true | id_part,id_True => true | id_part,id_Pair => true | id_part,id_Nil => true | id_part,id_match_4 => true | id_part,id_append => true | id_part,id_match_3 => true | id_part,id_part => 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 := (* test(x_0_,y_) -> True *) | R_xml_0_rule_0 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_True nil) (algebra.Alg.Term algebra.F.id_test ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil)) (* test(x_0_,y_) -> False *) | R_xml_0_rule_1 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_False nil) (algebra.Alg.Term algebra.F.id_test ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil)) (* append(l1_2_,l2_1_) -> match_0(l1_2_,l2_1_,l1_2_) *) | R_xml_0_rule_2 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_match_0 ((algebra.Alg.Var 3)::(algebra.Alg.Var 4):: (algebra.Alg.Var 3)::nil)) (algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Var 3):: (algebra.Alg.Var 4)::nil)) (* match_0(l1_2_,l2_1_,Nil) -> l2_1_ *) | R_xml_0_rule_3 : R_xml_0_rules (algebra.Alg.Var 4) (algebra.Alg.Term algebra.F.id_match_0 ((algebra.Alg.Var 3):: (algebra.Alg.Var 4)::(algebra.Alg.Term algebra.F.id_Nil nil)::nil)) (* match_0(l1_2_,l2_1_,Cons(x_,l_)) -> Cons(x_,append(l_,l2_1_)) *) | R_xml_0_rule_4 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_Cons ((algebra.Alg.Var 5):: (algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Var 6)::(algebra.Alg.Var 4)::nil))::nil)) (algebra.Alg.Term algebra.F.id_match_0 ((algebra.Alg.Var 3):: (algebra.Alg.Var 4)::(algebra.Alg.Term algebra.F.id_Cons ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* part(a_4_,l_3_) -> match_1(a_4_,l_3_,l_3_) *) | R_xml_0_rule_5 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_match_1 ((algebra.Alg.Var 7)::(algebra.Alg.Var 8):: (algebra.Alg.Var 8)::nil)) (algebra.Alg.Term algebra.F.id_part ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)) (* match_1(a_4_,l_3_,Nil) -> Pair(Nil,Nil) *) | R_xml_0_rule_6 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_Pair ((algebra.Alg.Term algebra.F.id_Nil nil)::(algebra.Alg.Term algebra.F.id_Nil nil)::nil)) (algebra.Alg.Term algebra.F.id_match_1 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::(algebra.Alg.Term algebra.F.id_Nil nil)::nil)) (* match_1(a_4_,l_3_,Cons(x_,l'_)) -> match_2(x_,l'_,a_4_,l_3_,part(a_4_,l'_)) *) | R_xml_0_rule_7 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_match_2 ((algebra.Alg.Var 5)::(algebra.Alg.Var 9):: (algebra.Alg.Var 7)::(algebra.Alg.Var 8):: (algebra.Alg.Term algebra.F.id_part ((algebra.Alg.Var 7):: (algebra.Alg.Var 9)::nil))::nil)) (algebra.Alg.Term algebra.F.id_match_1 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::(algebra.Alg.Term algebra.F.id_Cons ((algebra.Alg.Var 5)::(algebra.Alg.Var 9)::nil))::nil)) (* match_2(x_,l'_,a_4_,l_3_,Pair(l1_,l2_)) -> match_3(l1_,l2_,x_,l'_,a_4_,l_3_,test(a_4_,x_)) *) | R_xml_0_rule_8 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_match_3 ((algebra.Alg.Var 10)::(algebra.Alg.Var 11):: (algebra.Alg.Var 5)::(algebra.Alg.Var 9):: (algebra.Alg.Var 7)::(algebra.Alg.Var 8):: (algebra.Alg.Term algebra.F.id_test ((algebra.Alg.Var 7):: (algebra.Alg.Var 5)::nil))::nil)) (algebra.Alg.Term algebra.F.id_match_2 ((algebra.Alg.Var 5):: (algebra.Alg.Var 9)::(algebra.Alg.Var 7)::(algebra.Alg.Var 8):: (algebra.Alg.Term algebra.F.id_Pair ((algebra.Alg.Var 10):: (algebra.Alg.Var 11)::nil))::nil)) (* match_3(l1_,l2_,x_,l'_,a_4_,l_3_,False) -> Pair(Cons(x_,l1_),l2_) *) | R_xml_0_rule_9 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_Pair ((algebra.Alg.Term algebra.F.id_Cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 10)::nil))::(algebra.Alg.Var 11)::nil)) (algebra.Alg.Term algebra.F.id_match_3 ((algebra.Alg.Var 10):: (algebra.Alg.Var 11)::(algebra.Alg.Var 5)::(algebra.Alg.Var 9):: (algebra.Alg.Var 7)::(algebra.Alg.Var 8)::(algebra.Alg.Term algebra.F.id_False nil)::nil)) (* match_3(l1_,l2_,x_,l'_,a_4_,l_3_,True) -> Pair(l1_,Cons(x_,l2_)) *) | R_xml_0_rule_10 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_Pair ((algebra.Alg.Var 10):: (algebra.Alg.Term algebra.F.id_Cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 11)::nil))::nil)) (algebra.Alg.Term algebra.F.id_match_3 ((algebra.Alg.Var 10):: (algebra.Alg.Var 11)::(algebra.Alg.Var 5)::(algebra.Alg.Var 9):: (algebra.Alg.Var 7)::(algebra.Alg.Var 8)::(algebra.Alg.Term algebra.F.id_True nil)::nil)) (* quick(l_5_) -> match_4(l_5_,l_5_) *) | R_xml_0_rule_11 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_match_4 ((algebra.Alg.Var 12)::(algebra.Alg.Var 12)::nil)) (algebra.Alg.Term algebra.F.id_quick ((algebra.Alg.Var 12)::nil)) (* match_4(l_5_,Nil) -> Nil *) | R_xml_0_rule_12 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_Nil nil) (algebra.Alg.Term algebra.F.id_match_4 ((algebra.Alg.Var 12):: (algebra.Alg.Term algebra.F.id_Nil nil)::nil)) (* match_4(l_5_,Cons(a_,l'_)) -> match_5(a_,l'_,l_5_,part(a_,l'_)) *) | R_xml_0_rule_13 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_match_5 ((algebra.Alg.Var 13)::(algebra.Alg.Var 9):: (algebra.Alg.Var 12)::(algebra.Alg.Term algebra.F.id_part ((algebra.Alg.Var 13)::(algebra.Alg.Var 9)::nil))::nil)) (algebra.Alg.Term algebra.F.id_match_4 ((algebra.Alg.Var 12):: (algebra.Alg.Term algebra.F.id_Cons ((algebra.Alg.Var 13):: (algebra.Alg.Var 9)::nil))::nil)) (* match_5(a_,l'_,l_5_,Pair(l1_,l2_)) -> append(quick(l1_),Cons(a_,quick(l2_))) *) | R_xml_0_rule_14 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Term algebra.F.id_quick ((algebra.Alg.Var 10)::nil)):: (algebra.Alg.Term algebra.F.id_Cons ((algebra.Alg.Var 13)::(algebra.Alg.Term algebra.F.id_quick ((algebra.Alg.Var 11)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_match_5 ((algebra.Alg.Var 13):: (algebra.Alg.Var 9)::(algebra.Alg.Var 12)::(algebra.Alg.Term algebra.F.id_Pair ((algebra.Alg.Var 10):: (algebra.Alg.Var 11)::nil))::nil)) . Definition R_xml_0_rule_as_list_0 := ((algebra.Alg.Term algebra.F.id_test ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil)),(algebra.Alg.Term algebra.F.id_True nil)):: nil. Definition R_xml_0_rule_as_list_1 := ((algebra.Alg.Term algebra.F.id_test ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil)),(algebra.Alg.Term algebra.F.id_False nil)):: R_xml_0_rule_as_list_0. Definition R_xml_0_rule_as_list_2 := ((algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Var 3):: (algebra.Alg.Var 4)::nil)), (algebra.Alg.Term algebra.F.id_match_0 ((algebra.Alg.Var 3):: (algebra.Alg.Var 4)::(algebra.Alg.Var 3)::nil)))::R_xml_0_rule_as_list_1 . Definition R_xml_0_rule_as_list_3 := ((algebra.Alg.Term algebra.F.id_match_0 ((algebra.Alg.Var 3):: (algebra.Alg.Var 4)::(algebra.Alg.Term algebra.F.id_Nil nil)::nil)), (algebra.Alg.Var 4))::R_xml_0_rule_as_list_2. Definition R_xml_0_rule_as_list_4 := ((algebra.Alg.Term algebra.F.id_match_0 ((algebra.Alg.Var 3):: (algebra.Alg.Var 4)::(algebra.Alg.Term algebra.F.id_Cons ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Term algebra.F.id_Cons ((algebra.Alg.Var 5):: (algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Var 6):: (algebra.Alg.Var 4)::nil))::nil)))::R_xml_0_rule_as_list_3. Definition R_xml_0_rule_as_list_5 := ((algebra.Alg.Term algebra.F.id_part ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::nil)), (algebra.Alg.Term algebra.F.id_match_1 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::(algebra.Alg.Var 8)::nil)))::R_xml_0_rule_as_list_4 . Definition R_xml_0_rule_as_list_6 := ((algebra.Alg.Term algebra.F.id_match_1 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::(algebra.Alg.Term algebra.F.id_Nil nil)::nil)), (algebra.Alg.Term algebra.F.id_Pair ((algebra.Alg.Term algebra.F.id_Nil nil)::(algebra.Alg.Term algebra.F.id_Nil nil)::nil))):: R_xml_0_rule_as_list_5. Definition R_xml_0_rule_as_list_7 := ((algebra.Alg.Term algebra.F.id_match_1 ((algebra.Alg.Var 7):: (algebra.Alg.Var 8)::(algebra.Alg.Term algebra.F.id_Cons ((algebra.Alg.Var 5)::(algebra.Alg.Var 9)::nil))::nil)), (algebra.Alg.Term algebra.F.id_match_2 ((algebra.Alg.Var 5):: (algebra.Alg.Var 9)::(algebra.Alg.Var 7)::(algebra.Alg.Var 8):: (algebra.Alg.Term algebra.F.id_part ((algebra.Alg.Var 7):: (algebra.Alg.Var 9)::nil))::nil)))::R_xml_0_rule_as_list_6. Definition R_xml_0_rule_as_list_8 := ((algebra.Alg.Term algebra.F.id_match_2 ((algebra.Alg.Var 5):: (algebra.Alg.Var 9)::(algebra.Alg.Var 7)::(algebra.Alg.Var 8):: (algebra.Alg.Term algebra.F.id_Pair ((algebra.Alg.Var 10):: (algebra.Alg.Var 11)::nil))::nil)), (algebra.Alg.Term algebra.F.id_match_3 ((algebra.Alg.Var 10):: (algebra.Alg.Var 11)::(algebra.Alg.Var 5)::(algebra.Alg.Var 9):: (algebra.Alg.Var 7)::(algebra.Alg.Var 8)::(algebra.Alg.Term algebra.F.id_test ((algebra.Alg.Var 7):: (algebra.Alg.Var 5)::nil))::nil)))::R_xml_0_rule_as_list_7. Definition R_xml_0_rule_as_list_9 := ((algebra.Alg.Term algebra.F.id_match_3 ((algebra.Alg.Var 10):: (algebra.Alg.Var 11)::(algebra.Alg.Var 5)::(algebra.Alg.Var 9):: (algebra.Alg.Var 7)::(algebra.Alg.Var 8)::(algebra.Alg.Term algebra.F.id_False nil)::nil)), (algebra.Alg.Term algebra.F.id_Pair ((algebra.Alg.Term algebra.F.id_Cons ((algebra.Alg.Var 5)::(algebra.Alg.Var 10)::nil)):: (algebra.Alg.Var 11)::nil)))::R_xml_0_rule_as_list_8. Definition R_xml_0_rule_as_list_10 := ((algebra.Alg.Term algebra.F.id_match_3 ((algebra.Alg.Var 10):: (algebra.Alg.Var 11)::(algebra.Alg.Var 5)::(algebra.Alg.Var 9):: (algebra.Alg.Var 7)::(algebra.Alg.Var 8)::(algebra.Alg.Term algebra.F.id_True nil)::nil)), (algebra.Alg.Term algebra.F.id_Pair ((algebra.Alg.Var 10):: (algebra.Alg.Term algebra.F.id_Cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 11)::nil))::nil)))::R_xml_0_rule_as_list_9. Definition R_xml_0_rule_as_list_11 := ((algebra.Alg.Term algebra.F.id_quick ((algebra.Alg.Var 12)::nil)), (algebra.Alg.Term algebra.F.id_match_4 ((algebra.Alg.Var 12):: (algebra.Alg.Var 12)::nil)))::R_xml_0_rule_as_list_10. Definition R_xml_0_rule_as_list_12 := ((algebra.Alg.Term algebra.F.id_match_4 ((algebra.Alg.Var 12):: (algebra.Alg.Term algebra.F.id_Nil nil)::nil)), (algebra.Alg.Term algebra.F.id_Nil nil))::R_xml_0_rule_as_list_11. Definition R_xml_0_rule_as_list_13 := ((algebra.Alg.Term algebra.F.id_match_4 ((algebra.Alg.Var 12):: (algebra.Alg.Term algebra.F.id_Cons ((algebra.Alg.Var 13):: (algebra.Alg.Var 9)::nil))::nil)), (algebra.Alg.Term algebra.F.id_match_5 ((algebra.Alg.Var 13):: (algebra.Alg.Var 9)::(algebra.Alg.Var 12)::(algebra.Alg.Term algebra.F.id_part ((algebra.Alg.Var 13):: (algebra.Alg.Var 9)::nil))::nil)))::R_xml_0_rule_as_list_12. Definition R_xml_0_rule_as_list_14 := ((algebra.Alg.Term algebra.F.id_match_5 ((algebra.Alg.Var 13):: (algebra.Alg.Var 9)::(algebra.Alg.Var 12)::(algebra.Alg.Term algebra.F.id_Pair ((algebra.Alg.Var 10):: (algebra.Alg.Var 11)::nil))::nil)), (algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Term algebra.F.id_quick ((algebra.Alg.Var 10)::nil))::(algebra.Alg.Term algebra.F.id_Cons ((algebra.Alg.Var 13)::(algebra.Alg.Term algebra.F.id_quick ((algebra.Alg.Var 11)::nil))::nil))::nil))):: R_xml_0_rule_as_list_13. Definition R_xml_0_rule_as_list := R_xml_0_rule_as_list_14. 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.or16_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 15. injection H;intros ;subst;constructor 14. injection H;intros ;subst;constructor 13. injection H;intros ;subst;constructor 12. injection H;intros ;subst;constructor 11. injection H;intros ;subst;constructor 10. injection H;intros ;subst;constructor 9. injection H;intros ;subst;constructor 8. injection H;intros ;subst;constructor 7. injection H;intros ;subst;constructor 6. injection H;intros ;subst;constructor 5. injection H;intros ;subst;constructor 4. injection H;intros ;subst;constructor 3. injection H;intros ;subst;constructor 2. injection H;intros ;subst;constructor 1. elim H. Qed. Lemma R_xml_0_non_var : forall x t, ~R_xml_0_rules t (algebra.EQT.T.Var x). Proof. intros x t H. inversion H. Qed. Lemma R_xml_0_reg : forall s t, (R_xml_0_rules s t) -> forall x, In x (algebra.Alg.var_list s) ->In x (algebra.Alg.var_list t). Proof. intros s t H. inversion H;intros x Hx; (apply (more_list.mem_impl_in (@eq algebra.Alg.variable));[tauto|idtac]); apply (more_list.in_impl_mem (@eq algebra.Alg.variable)) in Hx; vm_compute in Hx|-*;tauto. Qed. Inductive and_5 (x15 x16 x17 x18 x19:Prop) : Prop := | conj_5 : x15->x16->x17->x18->x19->and_5 x15 x16 x17 x18 x19 . 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_5 (t = (algebra.Alg.Term algebra.F.id_False nil) -> t' = (algebra.Alg.Term algebra.F.id_False nil)) (forall x16 x18, t = (algebra.Alg.Term algebra.F.id_Cons (x16::x18::nil)) -> exists x15, exists x17, t' = (algebra.Alg.Term algebra.F.id_Cons (x15::x17::nil))/\ (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)) (t = (algebra.Alg.Term algebra.F.id_True nil) -> t' = (algebra.Alg.Term algebra.F.id_True nil)) (forall x16 x18, t = (algebra.Alg.Term algebra.F.id_Pair (x16::x18::nil)) -> exists x15, exists x17, t' = (algebra.Alg.Term algebra.F.id_Pair (x15::x17::nil))/\ (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)) (t = (algebra.Alg.Term algebra.F.id_Nil nil) -> t' = (algebra.Alg.Term algebra.F.id_Nil nil)). Proof. intros t t' H. induction H as [|y IH z z_to_y] using closure_extension.refl_trans_clos_ind2. constructor 1. intros H;intuition;constructor 1. intros x16 x18 H;exists x16;exists x18;intuition;constructor 1. intros H;intuition;constructor 1. intros x16 x18 H;exists x16;exists x18;intuition;constructor 1. intros H;intuition;constructor 1. inversion z_to_y as [t1 t2 H H0 H1|f l1 l2 H0 H H2];clear z_to_y;subst. inversion H as [t1 t2 sigma H2 H1 H0];clear H IH;subst;inversion H2; clear ;constructor;try (intros until 0 );clear ;intros abs; discriminate abs. destruct IH as [H_id_False H_id_Cons H_id_True H_id_Pair H_id_Nil]. constructor. clear H_id_Cons H_id_True H_id_Pair H_id_Nil;intros H;injection H; clear H;intros ;subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_False H_id_True H_id_Pair H_id_Nil;intros 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 x16 |- _ => destruct (H_id_Cons y x18 (refl_equal _)) as [x15 [x17]];intros ; intuition;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_Cons x16 y (refl_equal _)) as [x15 [x17]];intros ; intuition;exists x15;exists x17;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_False H_id_Cons H_id_Pair H_id_Nil;intros H;injection H; clear H;intros ;subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_False H_id_Cons H_id_True H_id_Nil;intros 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 x16 |- _ => destruct (H_id_Pair y x18 (refl_equal _)) as [x15 [x17]];intros ; intuition;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_Pair x16 y (refl_equal _)) as [x15 [x17]];intros ; intuition;exists x15;exists x17;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_False H_id_Cons H_id_True H_id_Pair;intros H;injection H; clear H;intros ;subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). Qed. Lemma id_False_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_False nil) -> t' = (algebra.Alg.Term algebra.F.id_False nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_Cons_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x16 x18, t = (algebra.Alg.Term algebra.F.id_Cons (x16::x18::nil)) -> exists x15, exists x17, t' = (algebra.Alg.Term algebra.F.id_Cons (x15::x17::nil))/\ (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_True_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_True nil) -> t' = (algebra.Alg.Term algebra.F.id_True nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_Pair_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x16 x18, t = (algebra.Alg.Term algebra.F.id_Pair (x16::x18::nil)) -> exists x15, exists x17, t' = (algebra.Alg.Term algebra.F.id_Pair (x15::x17::nil))/\ (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_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. 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_False nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_False_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_Cons (?x16::?x15::nil)) |- _ => let x16 := fresh "x" in (let x15 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_Cons_is_R_xml_0_constructor H (refl_equal _)) as [x16 [x15 [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_True nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_True_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_Pair (?x16::?x15::nil)) |- _ => let x16 := fresh "x" in (let x15 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_Pair_is_R_xml_0_constructor H (refl_equal _)) as [x16 [x15 [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 )) 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_False nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_False_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_Cons (?x16::?x15::nil)) |- _ => let x16 := fresh "x" in (let x15 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_Cons_is_R_xml_0_constructor H (refl_equal _)) as [x16 [x15 [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_True nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_True_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_Pair (?x16::?x15::nil)) |- _ => let x16 := fresh "x" in (let x15 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_Pair_is_R_xml_0_constructor H (refl_equal _)) as [x16 [x15 [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 ))) 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_test : A ->A ->A. Hypothesis P_id_match_1 : A ->A ->A ->A. Hypothesis P_id_match_0 : A ->A ->A ->A. Hypothesis P_id_quick : A ->A. Hypothesis P_id_False : A. Hypothesis P_id_match_2 : A ->A ->A ->A ->A ->A. Hypothesis P_id_Cons : A ->A ->A. Hypothesis P_id_match_5 : A ->A ->A ->A ->A. Hypothesis P_id_True : A. Hypothesis P_id_Pair : A ->A ->A. Hypothesis P_id_Nil : A. Hypothesis P_id_match_4 : A ->A ->A. Hypothesis P_id_append : A ->A ->A. Hypothesis P_id_match_3 : A ->A ->A ->A ->A ->A ->A ->A. Hypothesis P_id_part : A ->A ->A. Hypothesis P_id_test_monotonic : forall x16 x18 x17 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) ->P_id_test x16 x18 <= P_id_test x15 x17. Hypothesis P_id_match_1_monotonic : forall x16 x20 x18 x17 x19 x15, (A0 <= x20)/\ (x20 <= x19) -> (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> P_id_match_1 x16 x18 x20 <= P_id_match_1 x15 x17 x19. Hypothesis P_id_match_0_monotonic : forall x16 x20 x18 x17 x19 x15, (A0 <= x20)/\ (x20 <= x19) -> (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> P_id_match_0 x16 x18 x20 <= P_id_match_0 x15 x17 x19. Hypothesis P_id_quick_monotonic : forall x16 x15, (A0 <= x16)/\ (x16 <= x15) ->P_id_quick x16 <= P_id_quick x15. Hypothesis P_id_match_2_monotonic : forall x16 x24 x20 x18 x22 x17 x21 x19 x23 x15, (A0 <= x24)/\ (x24 <= x23) -> (A0 <= x22)/\ (x22 <= x21) -> (A0 <= x20)/\ (x20 <= x19) -> (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> P_id_match_2 x16 x18 x20 x22 x24 <= P_id_match_2 x15 x17 x19 x21 x23. Hypothesis P_id_Cons_monotonic : forall x16 x18 x17 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) ->P_id_Cons x16 x18 <= P_id_Cons x15 x17. Hypothesis P_id_match_5_monotonic : forall x16 x20 x18 x22 x17 x21 x19 x15, (A0 <= x22)/\ (x22 <= x21) -> (A0 <= x20)/\ (x20 <= x19) -> (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> P_id_match_5 x16 x18 x20 x22 <= P_id_match_5 x15 x17 x19 x21. Hypothesis P_id_Pair_monotonic : forall x16 x18 x17 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) ->P_id_Pair x16 x18 <= P_id_Pair x15 x17. Hypothesis P_id_match_4_monotonic : forall x16 x18 x17 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> P_id_match_4 x16 x18 <= P_id_match_4 x15 x17. Hypothesis P_id_append_monotonic : forall x16 x18 x17 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) ->P_id_append x16 x18 <= P_id_append x15 x17. Hypothesis P_id_match_3_monotonic : forall x16 x24 x20 x28 x18 x26 x22 x17 x25 x21 x19 x27 x23 x15, (A0 <= x28)/\ (x28 <= x27) -> (A0 <= x26)/\ (x26 <= x25) -> (A0 <= x24)/\ (x24 <= x23) -> (A0 <= x22)/\ (x22 <= x21) -> (A0 <= x20)/\ (x20 <= x19) -> (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> P_id_match_3 x16 x18 x20 x22 x24 x26 x28 <= P_id_match_3 x15 x17 x19 x21 x23 x25 x27. Hypothesis P_id_part_monotonic : forall x16 x18 x17 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) ->P_id_part x16 x18 <= P_id_part x15 x17. Hypothesis P_id_test_bounded : forall x16 x15, (A0 <= x15) ->(A0 <= x16) ->A0 <= P_id_test x16 x15. Hypothesis P_id_match_1_bounded : forall x16 x17 x15, (A0 <= x15) ->(A0 <= x16) ->(A0 <= x17) ->A0 <= P_id_match_1 x17 x16 x15. Hypothesis P_id_match_0_bounded : forall x16 x17 x15, (A0 <= x15) ->(A0 <= x16) ->(A0 <= x17) ->A0 <= P_id_match_0 x17 x16 x15. Hypothesis P_id_quick_bounded : forall x15, (A0 <= x15) ->A0 <= P_id_quick x15. Hypothesis P_id_False_bounded : A0 <= P_id_False . Hypothesis P_id_match_2_bounded : forall x16 x18 x17 x19 x15, (A0 <= x15) -> (A0 <= x16) -> (A0 <= x17) -> (A0 <= x18) ->(A0 <= x19) ->A0 <= P_id_match_2 x19 x18 x17 x16 x15. Hypothesis P_id_Cons_bounded : forall x16 x15, (A0 <= x15) ->(A0 <= x16) ->A0 <= P_id_Cons x16 x15. Hypothesis P_id_match_5_bounded : forall x16 x18 x17 x15, (A0 <= x15) -> (A0 <= x16) -> (A0 <= x17) ->(A0 <= x18) ->A0 <= P_id_match_5 x18 x17 x16 x15. Hypothesis P_id_True_bounded : A0 <= P_id_True . Hypothesis P_id_Pair_bounded : forall x16 x15, (A0 <= x15) ->(A0 <= x16) ->A0 <= P_id_Pair x16 x15. Hypothesis P_id_Nil_bounded : A0 <= P_id_Nil . Hypothesis P_id_match_4_bounded : forall x16 x15, (A0 <= x15) ->(A0 <= x16) ->A0 <= P_id_match_4 x16 x15. Hypothesis P_id_append_bounded : forall x16 x15, (A0 <= x15) ->(A0 <= x16) ->A0 <= P_id_append x16 x15. Hypothesis P_id_match_3_bounded : forall x16 x20 x18 x17 x21 x19 x15, (A0 <= x15) -> (A0 <= x16) -> (A0 <= x17) -> (A0 <= x18) -> (A0 <= x19) -> (A0 <= x20) -> (A0 <= x21) ->A0 <= P_id_match_3 x21 x20 x19 x18 x17 x16 x15. Hypothesis P_id_part_bounded : forall x16 x15, (A0 <= x15) ->(A0 <= x16) ->A0 <= P_id_part x16 x15. Fixpoint measure t { struct t } := match t with | (algebra.Alg.Term algebra.F.id_test (x16::x15::nil)) => P_id_test (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_1 (x17::x16::x15::nil)) => P_id_match_1 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_0 (x17::x16::x15::nil)) => P_id_match_0 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_quick (x15::nil)) => P_id_quick (measure x15) | (algebra.Alg.Term algebra.F.id_False nil) => P_id_False | (algebra.Alg.Term algebra.F.id_match_2 (x19::x18::x17::x16:: x15::nil)) => P_id_match_2 (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_Cons (x16::x15::nil)) => P_id_Cons (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16::x15::nil)) => P_id_match_5 (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_True nil) => P_id_True | (algebra.Alg.Term algebra.F.id_Pair (x16::x15::nil)) => P_id_Pair (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_Nil nil) => P_id_Nil | (algebra.Alg.Term algebra.F.id_match_4 (x16::x15::nil)) => P_id_match_4 (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_append (x16::x15::nil)) => P_id_append (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_3 (x21::x20::x19::x18::x17:: x16::x15::nil)) => P_id_match_3 (measure x21) (measure x20) (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_part (x16::x15::nil)) => P_id_part (measure x16) (measure x15) | _ => A0 end. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_test (x16::x15::nil)) => P_id_test (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_1 (x17::x16:: x15::nil)) => P_id_match_1 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_0 (x17::x16:: x15::nil)) => P_id_match_0 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_quick (x15::nil)) => P_id_quick (measure x15) | (algebra.Alg.Term algebra.F.id_False nil) => P_id_False | (algebra.Alg.Term algebra.F.id_match_2 (x19::x18::x17:: x16::x15::nil)) => P_id_match_2 (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_Cons (x16::x15::nil)) => P_id_Cons (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16:: x15::nil)) => P_id_match_5 (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_True nil) => P_id_True | (algebra.Alg.Term algebra.F.id_Pair (x16::x15::nil)) => P_id_Pair (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_Nil nil) => P_id_Nil | (algebra.Alg.Term algebra.F.id_match_4 (x16::x15::nil)) => P_id_match_4 (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_append (x16::x15::nil)) => P_id_append (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_3 (x21::x20::x19:: x18::x17::x16::x15::nil)) => P_id_match_3 (measure x21) (measure x20) (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_part (x16::x15::nil)) => P_id_part (measure x16) (measure x15) | _ => 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_test => P_id_test | algebra.F.id_match_1 => P_id_match_1 | algebra.F.id_match_0 => P_id_match_0 | algebra.F.id_quick => P_id_quick | algebra.F.id_False => P_id_False | algebra.F.id_match_2 => P_id_match_2 | algebra.F.id_Cons => P_id_Cons | algebra.F.id_match_5 => P_id_match_5 | algebra.F.id_True => P_id_True | algebra.F.id_Pair => P_id_Pair | algebra.F.id_Nil => P_id_Nil | algebra.F.id_match_4 => P_id_match_4 | algebra.F.id_append => P_id_append | algebra.F.id_match_3 => P_id_match_3 | algebra.F.id_part => P_id_part 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_test => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_match_1 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_match_0 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_quick => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_False => match l with | nil => _ | _::_ => _ end | algebra.F.id_match_2 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::nil => _ | _::_::_::_::_::_::_ => _ end | algebra.F.id_Cons => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_match_5 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::_ => _ end | algebra.F.id_True => match l with | nil => _ | _::_ => _ end | algebra.F.id_Pair => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_Nil => match l with | nil => _ | _::_ => _ end | algebra.F.id_match_4 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_append => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_match_3 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::nil => _ | _::_::_::_::_::_::nil => _ | _::_::_::_::_::_::_::nil => _ | _::_::_::_::_::_::_::_::_ => _ end | algebra.F.id_part => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end end;simpl in |-*;unfold eq_rect_r, eq_rect, sym_eq in |-*; try (reflexivity );f_equal ;auto. Qed. Lemma measure_bounded : forall t, A0 <= measure t. Proof. intros t. rewrite same_measure in |-*. apply (InterpGen.measure_bounded Aop). intros f. case f. vm_compute in |-*;intros ;apply P_id_test_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_1_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_0_bounded;assumption. vm_compute in |-*;intros ;apply P_id_quick_bounded;assumption. vm_compute in |-*;intros ;apply P_id_False_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_2_bounded;assumption. vm_compute in |-*;intros ;apply P_id_Cons_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_5_bounded;assumption. vm_compute in |-*;intros ;apply P_id_True_bounded;assumption. vm_compute in |-*;intros ;apply P_id_Pair_bounded;assumption. vm_compute in |-*;intros ;apply P_id_Nil_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_4_bounded;assumption. vm_compute in |-*;intros ;apply P_id_append_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_3_bounded;assumption. vm_compute in |-*;intros ;apply P_id_part_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_test_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_match_1_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_match_0_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_quick_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_match_2_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_Cons_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_match_5_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_Pair_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_match_4_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_append_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_match_3_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_part_monotonic;assumption. intros f. case f. vm_compute in |-*;intros ;apply P_id_test_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_1_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_0_bounded;assumption. vm_compute in |-*;intros ;apply P_id_quick_bounded;assumption. vm_compute in |-*;intros ;apply P_id_False_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_2_bounded;assumption. vm_compute in |-*;intros ;apply P_id_Cons_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_5_bounded;assumption. vm_compute in |-*;intros ;apply P_id_True_bounded;assumption. vm_compute in |-*;intros ;apply P_id_Pair_bounded;assumption. vm_compute in |-*;intros ;apply P_id_Nil_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_4_bounded;assumption. vm_compute in |-*;intros ;apply P_id_append_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_3_bounded;assumption. vm_compute in |-*;intros ;apply P_id_part_bounded;assumption. intros . do 2 (rewrite <- same_measure in |-*). apply rules_monotonic;assumption. assumption. Qed. Hypothesis P_id_MATCH_0 : A ->A ->A ->A. Hypothesis P_id_MATCH_5 : A ->A ->A ->A ->A. Hypothesis P_id_MATCH_3 : A ->A ->A ->A ->A ->A ->A ->A. Hypothesis P_id_MATCH_1 : A ->A ->A ->A. Hypothesis P_id_QUICK : A ->A. Hypothesis P_id_PART : A ->A ->A. Hypothesis P_id_TEST : A ->A ->A. Hypothesis P_id_MATCH_2 : A ->A ->A ->A ->A ->A. Hypothesis P_id_MATCH_4 : A ->A ->A. Hypothesis P_id_APPEND : A ->A ->A. Hypothesis P_id_MATCH_0_monotonic : forall x16 x20 x18 x17 x19 x15, (A0 <= x20)/\ (x20 <= x19) -> (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_0 x16 x18 x20 <= P_id_MATCH_0 x15 x17 x19. Hypothesis P_id_MATCH_5_monotonic : forall x16 x20 x18 x22 x17 x21 x19 x15, (A0 <= x22)/\ (x22 <= x21) -> (A0 <= x20)/\ (x20 <= x19) -> (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_5 x16 x18 x20 x22 <= P_id_MATCH_5 x15 x17 x19 x21. Hypothesis P_id_MATCH_3_monotonic : forall x16 x24 x20 x28 x18 x26 x22 x17 x25 x21 x19 x27 x23 x15, (A0 <= x28)/\ (x28 <= x27) -> (A0 <= x26)/\ (x26 <= x25) -> (A0 <= x24)/\ (x24 <= x23) -> (A0 <= x22)/\ (x22 <= x21) -> (A0 <= x20)/\ (x20 <= x19) -> (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_3 x16 x18 x20 x22 x24 x26 x28 <= P_id_MATCH_3 x15 x17 x19 x21 x23 x25 x27. Hypothesis P_id_MATCH_1_monotonic : forall x16 x20 x18 x17 x19 x15, (A0 <= x20)/\ (x20 <= x19) -> (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_1 x16 x18 x20 <= P_id_MATCH_1 x15 x17 x19. Hypothesis P_id_QUICK_monotonic : forall x16 x15, (A0 <= x16)/\ (x16 <= x15) ->P_id_QUICK x16 <= P_id_QUICK x15. Hypothesis P_id_PART_monotonic : forall x16 x18 x17 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) ->P_id_PART x16 x18 <= P_id_PART x15 x17. Hypothesis P_id_TEST_monotonic : forall x16 x18 x17 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) ->P_id_TEST x16 x18 <= P_id_TEST x15 x17. Hypothesis P_id_MATCH_2_monotonic : forall x16 x24 x20 x18 x22 x17 x21 x19 x23 x15, (A0 <= x24)/\ (x24 <= x23) -> (A0 <= x22)/\ (x22 <= x21) -> (A0 <= x20)/\ (x20 <= x19) -> (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_2 x16 x18 x20 x22 x24 <= P_id_MATCH_2 x15 x17 x19 x21 x23. Hypothesis P_id_MATCH_4_monotonic : forall x16 x18 x17 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_4 x16 x18 <= P_id_MATCH_4 x15 x17. Hypothesis P_id_APPEND_monotonic : forall x16 x18 x17 x15, (A0 <= x18)/\ (x18 <= x17) -> (A0 <= x16)/\ (x16 <= x15) ->P_id_APPEND x16 x18 <= P_id_APPEND x15 x17. Definition marked_measure t := match t with | (algebra.Alg.Term algebra.F.id_match_0 (x17::x16::x15::nil)) => P_id_MATCH_0 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16::x15::nil)) => P_id_MATCH_5 (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_3 (x21::x20::x19::x18::x17:: x16::x15::nil)) => P_id_MATCH_3 (measure x21) (measure x20) (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_1 (x17::x16::x15::nil)) => P_id_MATCH_1 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_quick (x15::nil)) => P_id_QUICK (measure x15) | (algebra.Alg.Term algebra.F.id_part (x16::x15::nil)) => P_id_PART (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_test (x16::x15::nil)) => P_id_TEST (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_2 (x19::x18::x17::x16:: x15::nil)) => P_id_MATCH_2 (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_4 (x16::x15::nil)) => P_id_MATCH_4 (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_append (x16::x15::nil)) => P_id_APPEND (measure x16) (measure x15) | _ => 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 x16 x15;apply (P_id_PART x16 x15). 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 x21 x20 x19 x18 x17 x16 x15; apply (P_id_MATCH_3 x21 x20 x19 x18 x17 x16 x15). 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 x16 x15;apply (P_id_APPEND x16 x15). 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 x16 x15;apply (P_id_MATCH_4 x16 x15). 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 x18 x17 x16 x15; apply (P_id_MATCH_5 x18 x17 x16 x15). 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 x19 x18 x17 x16 x15; apply (P_id_MATCH_2 x19 x18 x17 x16 x15). 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;apply (P_id_QUICK x15). 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 x17 x16 x15;apply (P_id_MATCH_0 x17 x16 x15). 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 x17 x16 x15;apply (P_id_MATCH_1 x17 x16 x15). 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 x16 x15;apply (P_id_TEST x16 x15). 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_test => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_match_1 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_match_0 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_quick => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_False => match l with | nil => _ | _::_ => _ end | algebra.F.id_match_2 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::nil => _ | _::_::_::_::_::_::_ => _ end | algebra.F.id_Cons => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_match_5 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::_ => _ end | algebra.F.id_True => match l with | nil => _ | _::_ => _ end | algebra.F.id_Pair => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_Nil => match l with | nil => _ | _::_ => _ end | algebra.F.id_match_4 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_append => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_match_3 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::nil => _ | _::_::_::_::_::_::nil => _ | _::_::_::_::_::_::_::nil => _ | _::_::_::_::_::_::_::_::_ => _ end | algebra.F.id_part => match l with | nil => _ | _::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 . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . 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 . 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 . Qed. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_match_0 (x17:: x16::x15::nil)) => P_id_MATCH_0 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_5 (x18:: x17::x16::x15::nil)) => P_id_MATCH_5 (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_3 (x21:: x20::x19::x18::x17::x16::x15::nil)) => P_id_MATCH_3 (measure x21) (measure x20) (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_1 (x17:: x16::x15::nil)) => P_id_MATCH_1 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_quick (x15::nil)) => P_id_QUICK (measure x15) | (algebra.Alg.Term algebra.F.id_part (x16:: x15::nil)) => P_id_PART (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_test (x16:: x15::nil)) => P_id_TEST (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_2 (x19:: x18::x17::x16::x15::nil)) => P_id_MATCH_2 (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_4 (x16:: x15::nil)) => P_id_MATCH_4 (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_append (x16:: x15::nil)) => P_id_APPEND (measure x16) (measure x15) | _ => 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_test_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_match_1_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_match_0_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_quick_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_match_2_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_Cons_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_match_5_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_Pair_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_match_4_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_append_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_match_3_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_part_monotonic;assumption. clear f. intros f. case f. vm_compute in |-*;intros ;apply P_id_test_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_1_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_0_bounded;assumption. vm_compute in |-*;intros ;apply P_id_quick_bounded;assumption. vm_compute in |-*;intros ;apply P_id_False_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_2_bounded;assumption. vm_compute in |-*;intros ;apply P_id_Cons_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_5_bounded;assumption. vm_compute in |-*;intros ;apply P_id_True_bounded;assumption. vm_compute in |-*;intros ;apply P_id_Pair_bounded;assumption. vm_compute in |-*;intros ;apply P_id_Nil_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_4_bounded;assumption. vm_compute in |-*;intros ;apply P_id_append_bounded;assumption. vm_compute in |-*;intros ;apply P_id_match_3_bounded;assumption. vm_compute in |-*;intros ;apply P_id_part_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_PART_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_MATCH_3_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_APPEND_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_MATCH_4_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_MATCH_5_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_MATCH_2_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_QUICK_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_MATCH_0_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_MATCH_1_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_TEST_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_test : Z ->Z ->Z. Hypothesis P_id_match_1 : Z ->Z ->Z ->Z. Hypothesis P_id_match_0 : Z ->Z ->Z ->Z. Hypothesis P_id_quick : Z ->Z. Hypothesis P_id_False : Z. Hypothesis P_id_match_2 : Z ->Z ->Z ->Z ->Z ->Z. Hypothesis P_id_Cons : Z ->Z ->Z. Hypothesis P_id_match_5 : Z ->Z ->Z ->Z ->Z. Hypothesis P_id_True : Z. Hypothesis P_id_Pair : Z ->Z ->Z. Hypothesis P_id_Nil : Z. Hypothesis P_id_match_4 : Z ->Z ->Z. Hypothesis P_id_append : Z ->Z ->Z. Hypothesis P_id_match_3 : Z ->Z ->Z ->Z ->Z ->Z ->Z ->Z. Hypothesis P_id_part : Z ->Z ->Z. Hypothesis P_id_test_monotonic : forall x16 x18 x17 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_test x16 x18 <= P_id_test x15 x17. Hypothesis P_id_match_1_monotonic : forall x16 x20 x18 x17 x19 x15, (min_value <= x20)/\ (x20 <= x19) -> (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_match_1 x16 x18 x20 <= P_id_match_1 x15 x17 x19. Hypothesis P_id_match_0_monotonic : forall x16 x20 x18 x17 x19 x15, (min_value <= x20)/\ (x20 <= x19) -> (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_match_0 x16 x18 x20 <= P_id_match_0 x15 x17 x19. Hypothesis P_id_quick_monotonic : forall x16 x15, (min_value <= x16)/\ (x16 <= x15) ->P_id_quick x16 <= P_id_quick x15. Hypothesis P_id_match_2_monotonic : forall x16 x24 x20 x18 x22 x17 x21 x19 x23 x15, (min_value <= x24)/\ (x24 <= x23) -> (min_value <= x22)/\ (x22 <= x21) -> (min_value <= x20)/\ (x20 <= x19) -> (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_match_2 x16 x18 x20 x22 x24 <= P_id_match_2 x15 x17 x19 x21 x23. Hypothesis P_id_Cons_monotonic : forall x16 x18 x17 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_Cons x16 x18 <= P_id_Cons x15 x17. Hypothesis P_id_match_5_monotonic : forall x16 x20 x18 x22 x17 x21 x19 x15, (min_value <= x22)/\ (x22 <= x21) -> (min_value <= x20)/\ (x20 <= x19) -> (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_match_5 x16 x18 x20 x22 <= P_id_match_5 x15 x17 x19 x21. Hypothesis P_id_Pair_monotonic : forall x16 x18 x17 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_Pair x16 x18 <= P_id_Pair x15 x17. Hypothesis P_id_match_4_monotonic : forall x16 x18 x17 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_match_4 x16 x18 <= P_id_match_4 x15 x17. Hypothesis P_id_append_monotonic : forall x16 x18 x17 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_append x16 x18 <= P_id_append x15 x17. Hypothesis P_id_match_3_monotonic : forall x16 x24 x20 x28 x18 x26 x22 x17 x25 x21 x19 x27 x23 x15, (min_value <= x28)/\ (x28 <= x27) -> (min_value <= x26)/\ (x26 <= x25) -> (min_value <= x24)/\ (x24 <= x23) -> (min_value <= x22)/\ (x22 <= x21) -> (min_value <= x20)/\ (x20 <= x19) -> (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_match_3 x16 x18 x20 x22 x24 x26 x28 <= P_id_match_3 x15 x17 x19 x21 x23 x25 x27. Hypothesis P_id_part_monotonic : forall x16 x18 x17 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_part x16 x18 <= P_id_part x15 x17. Hypothesis P_id_test_bounded : forall x16 x15, (min_value <= x15) ->(min_value <= x16) ->min_value <= P_id_test x16 x15. Hypothesis P_id_match_1_bounded : forall x16 x17 x15, (min_value <= x15) -> (min_value <= x16) -> (min_value <= x17) ->min_value <= P_id_match_1 x17 x16 x15. Hypothesis P_id_match_0_bounded : forall x16 x17 x15, (min_value <= x15) -> (min_value <= x16) -> (min_value <= x17) ->min_value <= P_id_match_0 x17 x16 x15. Hypothesis P_id_quick_bounded : forall x15, (min_value <= x15) ->min_value <= P_id_quick x15. Hypothesis P_id_False_bounded : min_value <= P_id_False . Hypothesis P_id_match_2_bounded : forall x16 x18 x17 x19 x15, (min_value <= x15) -> (min_value <= x16) -> (min_value <= x17) -> (min_value <= x18) -> (min_value <= x19) ->min_value <= P_id_match_2 x19 x18 x17 x16 x15. Hypothesis P_id_Cons_bounded : forall x16 x15, (min_value <= x15) ->(min_value <= x16) ->min_value <= P_id_Cons x16 x15. Hypothesis P_id_match_5_bounded : forall x16 x18 x17 x15, (min_value <= x15) -> (min_value <= x16) -> (min_value <= x17) -> (min_value <= x18) ->min_value <= P_id_match_5 x18 x17 x16 x15. Hypothesis P_id_True_bounded : min_value <= P_id_True . Hypothesis P_id_Pair_bounded : forall x16 x15, (min_value <= x15) ->(min_value <= x16) ->min_value <= P_id_Pair x16 x15. Hypothesis P_id_Nil_bounded : min_value <= P_id_Nil . Hypothesis P_id_match_4_bounded : forall x16 x15, (min_value <= x15) -> (min_value <= x16) ->min_value <= P_id_match_4 x16 x15. Hypothesis P_id_append_bounded : forall x16 x15, (min_value <= x15) -> (min_value <= x16) ->min_value <= P_id_append x16 x15. Hypothesis P_id_match_3_bounded : forall x16 x20 x18 x17 x21 x19 x15, (min_value <= x15) -> (min_value <= x16) -> (min_value <= x17) -> (min_value <= x18) -> (min_value <= x19) -> (min_value <= x20) -> (min_value <= x21) -> min_value <= P_id_match_3 x21 x20 x19 x18 x17 x16 x15. Hypothesis P_id_part_bounded : forall x16 x15, (min_value <= x15) ->(min_value <= x16) ->min_value <= P_id_part x16 x15. Definition measure := Interp.measure min_value P_id_test P_id_match_1 P_id_match_0 P_id_quick P_id_False P_id_match_2 P_id_Cons P_id_match_5 P_id_True P_id_Pair P_id_Nil P_id_match_4 P_id_append P_id_match_3 P_id_part. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_test (x16::x15::nil)) => P_id_test (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_1 (x17::x16:: x15::nil)) => P_id_match_1 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_0 (x17::x16:: x15::nil)) => P_id_match_0 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_quick (x15::nil)) => P_id_quick (measure x15) | (algebra.Alg.Term algebra.F.id_False nil) => P_id_False | (algebra.Alg.Term algebra.F.id_match_2 (x19::x18::x17:: x16::x15::nil)) => P_id_match_2 (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_Cons (x16::x15::nil)) => P_id_Cons (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16:: x15::nil)) => P_id_match_5 (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_True nil) => P_id_True | (algebra.Alg.Term algebra.F.id_Pair (x16::x15::nil)) => P_id_Pair (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_Nil nil) => P_id_Nil | (algebra.Alg.Term algebra.F.id_match_4 (x16::x15::nil)) => P_id_match_4 (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_append (x16::x15::nil)) => P_id_append (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_3 (x21::x20::x19:: x18::x17::x16::x15::nil)) => P_id_match_3 (measure x21) (measure x20) (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_part (x16::x15::nil)) => P_id_part (measure x16) (measure x15) | _ => 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_test_monotonic;assumption. intros ;apply P_id_match_1_monotonic;assumption. intros ;apply P_id_match_0_monotonic;assumption. intros ;apply P_id_quick_monotonic;assumption. intros ;apply P_id_match_2_monotonic;assumption. intros ;apply P_id_Cons_monotonic;assumption. intros ;apply P_id_match_5_monotonic;assumption. intros ;apply P_id_Pair_monotonic;assumption. intros ;apply P_id_match_4_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_match_3_monotonic;assumption. intros ;apply P_id_part_monotonic;assumption. intros ;apply P_id_test_bounded;assumption. intros ;apply P_id_match_1_bounded;assumption. intros ;apply P_id_match_0_bounded;assumption. intros ;apply P_id_quick_bounded;assumption. intros ;apply P_id_False_bounded;assumption. intros ;apply P_id_match_2_bounded;assumption. intros ;apply P_id_Cons_bounded;assumption. intros ;apply P_id_match_5_bounded;assumption. intros ;apply P_id_True_bounded;assumption. intros ;apply P_id_Pair_bounded;assumption. intros ;apply P_id_Nil_bounded;assumption. intros ;apply P_id_match_4_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_match_3_bounded;assumption. intros ;apply P_id_part_bounded;assumption. apply rules_monotonic. Qed. Hypothesis P_id_MATCH_0 : Z ->Z ->Z ->Z. Hypothesis P_id_MATCH_5 : Z ->Z ->Z ->Z ->Z. Hypothesis P_id_MATCH_3 : Z ->Z ->Z ->Z ->Z ->Z ->Z ->Z. Hypothesis P_id_MATCH_1 : Z ->Z ->Z ->Z. Hypothesis P_id_QUICK : Z ->Z. Hypothesis P_id_PART : Z ->Z ->Z. Hypothesis P_id_TEST : Z ->Z ->Z. Hypothesis P_id_MATCH_2 : Z ->Z ->Z ->Z ->Z ->Z. Hypothesis P_id_MATCH_4 : Z ->Z ->Z. Hypothesis P_id_APPEND : Z ->Z ->Z. Hypothesis P_id_MATCH_0_monotonic : forall x16 x20 x18 x17 x19 x15, (min_value <= x20)/\ (x20 <= x19) -> (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_MATCH_0 x16 x18 x20 <= P_id_MATCH_0 x15 x17 x19. Hypothesis P_id_MATCH_5_monotonic : forall x16 x20 x18 x22 x17 x21 x19 x15, (min_value <= x22)/\ (x22 <= x21) -> (min_value <= x20)/\ (x20 <= x19) -> (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_MATCH_5 x16 x18 x20 x22 <= P_id_MATCH_5 x15 x17 x19 x21. Hypothesis P_id_MATCH_3_monotonic : forall x16 x24 x20 x28 x18 x26 x22 x17 x25 x21 x19 x27 x23 x15, (min_value <= x28)/\ (x28 <= x27) -> (min_value <= x26)/\ (x26 <= x25) -> (min_value <= x24)/\ (x24 <= x23) -> (min_value <= x22)/\ (x22 <= x21) -> (min_value <= x20)/\ (x20 <= x19) -> (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_MATCH_3 x16 x18 x20 x22 x24 x26 x28 <= P_id_MATCH_3 x15 x17 x19 x21 x23 x25 x27. Hypothesis P_id_MATCH_1_monotonic : forall x16 x20 x18 x17 x19 x15, (min_value <= x20)/\ (x20 <= x19) -> (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_MATCH_1 x16 x18 x20 <= P_id_MATCH_1 x15 x17 x19. Hypothesis P_id_QUICK_monotonic : forall x16 x15, (min_value <= x16)/\ (x16 <= x15) ->P_id_QUICK x16 <= P_id_QUICK x15. Hypothesis P_id_PART_monotonic : forall x16 x18 x17 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_PART x16 x18 <= P_id_PART x15 x17. Hypothesis P_id_TEST_monotonic : forall x16 x18 x17 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_TEST x16 x18 <= P_id_TEST x15 x17. Hypothesis P_id_MATCH_2_monotonic : forall x16 x24 x20 x18 x22 x17 x21 x19 x23 x15, (min_value <= x24)/\ (x24 <= x23) -> (min_value <= x22)/\ (x22 <= x21) -> (min_value <= x20)/\ (x20 <= x19) -> (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_MATCH_2 x16 x18 x20 x22 x24 <= P_id_MATCH_2 x15 x17 x19 x21 x23. Hypothesis P_id_MATCH_4_monotonic : forall x16 x18 x17 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_MATCH_4 x16 x18 <= P_id_MATCH_4 x15 x17. Hypothesis P_id_APPEND_monotonic : forall x16 x18 x17 x15, (min_value <= x18)/\ (x18 <= x17) -> (min_value <= x16)/\ (x16 <= x15) -> P_id_APPEND x16 x18 <= P_id_APPEND x15 x17. Definition marked_measure := Interp.marked_measure min_value P_id_test P_id_match_1 P_id_match_0 P_id_quick P_id_False P_id_match_2 P_id_Cons P_id_match_5 P_id_True P_id_Pair P_id_Nil P_id_match_4 P_id_append P_id_match_3 P_id_part P_id_MATCH_0 P_id_MATCH_5 P_id_MATCH_3 P_id_MATCH_1 P_id_QUICK P_id_PART P_id_TEST P_id_MATCH_2 P_id_MATCH_4 P_id_APPEND. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_match_0 (x17:: x16::x15::nil)) => P_id_MATCH_0 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_5 (x18:: x17::x16::x15::nil)) => P_id_MATCH_5 (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_3 (x21:: x20::x19::x18::x17::x16::x15::nil)) => P_id_MATCH_3 (measure x21) (measure x20) (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_1 (x17:: x16::x15::nil)) => P_id_MATCH_1 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_quick (x15::nil)) => P_id_QUICK (measure x15) | (algebra.Alg.Term algebra.F.id_part (x16:: x15::nil)) => P_id_PART (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_test (x16:: x15::nil)) => P_id_TEST (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_2 (x19:: x18::x17::x16::x15::nil)) => P_id_MATCH_2 (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_4 (x16:: x15::nil)) => P_id_MATCH_4 (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_append (x16:: x15::nil)) => P_id_APPEND (measure x16) (measure x15) | _ => 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_test_monotonic;assumption. intros ;apply P_id_match_1_monotonic;assumption. intros ;apply P_id_match_0_monotonic;assumption. intros ;apply P_id_quick_monotonic;assumption. intros ;apply P_id_match_2_monotonic;assumption. intros ;apply P_id_Cons_monotonic;assumption. intros ;apply P_id_match_5_monotonic;assumption. intros ;apply P_id_Pair_monotonic;assumption. intros ;apply P_id_match_4_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_match_3_monotonic;assumption. intros ;apply P_id_part_monotonic;assumption. intros ;apply P_id_test_bounded;assumption. intros ;apply P_id_match_1_bounded;assumption. intros ;apply P_id_match_0_bounded;assumption. intros ;apply P_id_quick_bounded;assumption. intros ;apply P_id_False_bounded;assumption. intros ;apply P_id_match_2_bounded;assumption. intros ;apply P_id_Cons_bounded;assumption. intros ;apply P_id_match_5_bounded;assumption. intros ;apply P_id_True_bounded;assumption. intros ;apply P_id_Pair_bounded;assumption. intros ;apply P_id_Nil_bounded;assumption. intros ;apply P_id_match_4_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_match_3_bounded;assumption. intros ;apply P_id_part_bounded;assumption. apply rules_monotonic. intros ;apply P_id_MATCH_0_monotonic;assumption. intros ;apply P_id_MATCH_5_monotonic;assumption. intros ;apply P_id_MATCH_3_monotonic;assumption. intros ;apply P_id_MATCH_1_monotonic;assumption. intros ;apply P_id_QUICK_monotonic;assumption. intros ;apply P_id_PART_monotonic;assumption. intros ;apply P_id_TEST_monotonic;assumption. intros ;apply P_id_MATCH_2_monotonic;assumption. intros ;apply P_id_MATCH_4_monotonic;assumption. intros ;apply P_id_APPEND_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 x16 x4 x3 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x15) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_match_0 (x3::x4::x3::nil)) (algebra.Alg.Term algebra.F.id_append (x16::x15::nil)) (* *) | DP_R_xml_0_1 : forall x16 x4 x6 x17 x5 x3 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Cons (x5::x6::nil)) x15) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_append (x6::x4::nil)) (algebra.Alg.Term algebra.F.id_match_0 (x17::x16::x15::nil)) (* *) | DP_R_xml_0_2 : forall x16 x8 x7 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x7 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x8 x15) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_match_1 (x7::x8::x8::nil)) (algebra.Alg.Term algebra.F.id_part (x16::x15::nil)) (* *) | DP_R_xml_0_3 : forall x16 x8 x17 x9 x5 x7 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x7 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x8 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Cons (x5::x9::nil)) x15) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_match_2 (x5::x9::x7::x8:: (algebra.Alg.Term algebra.F.id_part (x7::x9::nil))::nil)) (algebra.Alg.Term algebra.F.id_match_1 (x17::x16::x15::nil)) (* *) | DP_R_xml_0_4 : forall x16 x8 x17 x9 x5 x7 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x7 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x8 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Cons (x5::x9::nil)) x15) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_part (x7::x9::nil)) (algebra.Alg.Term algebra.F.id_match_1 (x17::x16::x15::nil)) (* *) | DP_R_xml_0_5 : forall x16 x8 x18 x10 x17 x9 x5 x19 x11 x7 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x19) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x9 x18) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x7 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x8 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Pair (x10::x11::nil)) x15) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_match_3 (x10::x11::x5:: x9::x7::x8::(algebra.Alg.Term algebra.F.id_test (x7:: x5::nil))::nil)) (algebra.Alg.Term algebra.F.id_match_2 (x19::x18::x17::x16:: x15::nil)) (* *) | DP_R_xml_0_6 : forall x16 x8 x18 x10 x17 x9 x5 x19 x11 x7 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x19) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x9 x18) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x7 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x8 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Pair (x10::x11::nil)) x15) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_test (x7::x5::nil)) (algebra.Alg.Term algebra.F.id_match_2 (x19::x18::x17::x16:: x15::nil)) (* *) | DP_R_xml_0_7 : forall x12 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x15) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_match_4 (x12::x12::nil)) (algebra.Alg.Term algebra.F.id_quick (x15::nil)) (* *) | DP_R_xml_0_8 : forall x16 x12 x9 x13 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Cons (x13::x9::nil)) x15) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_match_5 (x13::x9::x12:: (algebra.Alg.Term algebra.F.id_part (x13::x9::nil))::nil)) (algebra.Alg.Term algebra.F.id_match_4 (x16::x15::nil)) (* *) | DP_R_xml_0_9 : forall x16 x12 x9 x13 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Cons (x13::x9::nil)) x15) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_part (x13::x9::nil)) (algebra.Alg.Term algebra.F.id_match_4 (x16::x15::nil)) (* *) | DP_R_xml_0_10 : forall x16 x12 x18 x10 x17 x9 x13 x11 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x13 x18) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x9 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Pair (x10::x11::nil)) x15) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Term algebra.F.id_quick (x10::nil))::(algebra.Alg.Term algebra.F.id_Cons (x13::(algebra.Alg.Term algebra.F.id_quick (x11::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16::x15::nil)) (* *) | DP_R_xml_0_11 : forall x16 x12 x18 x10 x17 x9 x13 x11 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x13 x18) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x9 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Pair (x10::x11::nil)) x15) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_quick (x10::nil)) (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16::x15::nil)) (* *) | DP_R_xml_0_12 : forall x16 x12 x18 x10 x17 x9 x13 x11 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x13 x18) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x9 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Pair (x10::x11::nil)) x15) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_quick (x11::nil)) (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16::x15::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 x16 x8 x18 x10 x17 x9 x5 x19 x11 x7 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x19) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x9 x18) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x7 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x8 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Pair (x10::x11::nil)) x15) -> DP_R_xml_0_non_scc_1 (algebra.Alg.Term algebra.F.id_test (x7:: x5::nil)) (algebra.Alg.Term algebra.F.id_match_2 (x19::x18::x17::x16:: x15::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 x16 x8 x18 x10 x17 x9 x5 x19 x11 x7 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x5 x19) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x9 x18) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x7 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x8 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Pair (x10::x11::nil)) x15) -> DP_R_xml_0_non_scc_2 (algebra.Alg.Term algebra.F.id_match_3 (x10:: x11::x5::x9::x7::x8::(algebra.Alg.Term algebra.F.id_test (x7::x5::nil))::nil)) (algebra.Alg.Term algebra.F.id_match_2 (x19::x18::x17::x16:: x15::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 x16 x8 x17 x9 x5 x7 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x7 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x8 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Cons (x5::x9::nil)) x15) -> DP_R_xml_0_non_scc_3 (algebra.Alg.Term algebra.F.id_match_2 (x5:: x9::x7::x8::(algebra.Alg.Term algebra.F.id_part (x7::x9::nil))::nil)) (algebra.Alg.Term algebra.F.id_match_1 (x17::x16::x15::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; (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' )))). Qed. Inductive DP_R_xml_0_scc_4 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_4_0 : forall x16 x8 x17 x9 x5 x7 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x7 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x8 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Cons (x5::x9::nil)) x15) -> DP_R_xml_0_scc_4 (algebra.Alg.Term algebra.F.id_part (x7::x9::nil)) (algebra.Alg.Term algebra.F.id_match_1 (x17::x16::x15::nil)) (* *) | DP_R_xml_0_scc_4_1 : forall x16 x8 x7 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x7 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x8 x15) -> DP_R_xml_0_scc_4 (algebra.Alg.Term algebra.F.id_match_1 (x7::x8:: x8::nil)) (algebra.Alg.Term algebra.F.id_part (x16::x15::nil)) . Module WF_DP_R_xml_0_scc_4. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_test (x15:Z) (x16:Z) := 1. Definition P_id_match_1 (x15:Z) (x16:Z) (x17:Z) := 1 + 1* x17. Definition P_id_match_0 (x15:Z) (x16:Z) (x17:Z) := 1* x16 + 1* x17. Definition P_id_quick (x15:Z) := 1* x15. Definition P_id_False := 1. Definition P_id_match_2 (x15:Z) (x16:Z) (x17:Z) (x18:Z) (x19:Z) := 2 + 1* x19. Definition P_id_Cons (x15:Z) (x16:Z) := 2 + 1* x16. Definition P_id_match_5 (x15:Z) (x16:Z) (x17:Z) (x18:Z) := 1 + 1* x18. Definition P_id_True := 1. Definition P_id_Pair (x15:Z) (x16:Z) := 1 + 1* x15 + 1* x16. Definition P_id_Nil := 0. Definition P_id_match_4 (x15:Z) (x16:Z) := 1* x16. Definition P_id_append (x15:Z) (x16:Z) := 1* x15 + 1* x16. Definition P_id_match_3 (x15:Z) (x16:Z) (x17:Z) (x18:Z) (x19:Z) (x20: Z) (x21:Z) := 1 + 1* x15 + 1* x16 + 2* x21. Definition P_id_part (x15:Z) (x16:Z) := 1 + 1* x16. Lemma P_id_test_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_test x16 x18 <= P_id_test x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_1_monotonic : forall x16 x20 x18 x17 x19 x15, (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_1 x16 x18 x20 <= P_id_match_1 x15 x17 x19. Proof. intros x20 x19 x18 x17 x16 x15. 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_match_0_monotonic : forall x16 x20 x18 x17 x19 x15, (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_0 x16 x18 x20 <= P_id_match_0 x15 x17 x19. Proof. intros x20 x19 x18 x17 x16 x15. 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_quick_monotonic : forall x16 x15, (0 <= x16)/\ (x16 <= x15) ->P_id_quick x16 <= P_id_quick x15. Proof. intros x16 x15. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_2_monotonic : forall x16 x24 x20 x18 x22 x17 x21 x19 x23 x15, (0 <= x24)/\ (x24 <= x23) -> (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_2 x16 x18 x20 x22 x24 <= P_id_match_2 x15 x17 x19 x21 x23. Proof. intros x24 x23 x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Cons_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_Cons x16 x18 <= P_id_Cons x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_5_monotonic : forall x16 x20 x18 x22 x17 x21 x19 x15, (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_5 x16 x18 x20 x22 <= P_id_match_5 x15 x17 x19 x21. Proof. intros x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Pair_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_Pair x16 x18 <= P_id_Pair x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_4_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_4 x16 x18 <= P_id_match_4 x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_append x16 x18 <= P_id_append x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_3_monotonic : forall x16 x24 x20 x28 x18 x26 x22 x17 x25 x21 x19 x27 x23 x15, (0 <= x28)/\ (x28 <= x27) -> (0 <= x26)/\ (x26 <= x25) -> (0 <= x24)/\ (x24 <= x23) -> (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_3 x16 x18 x20 x22 x24 x26 x28 <= P_id_match_3 x15 x17 x19 x21 x23 x25 x27. Proof. intros x28 x27 x26 x25 x24 x23 x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. intros [H_13 H_12]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_part_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_part x16 x18 <= P_id_part x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_test_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_test x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_1_bounded : forall x16 x17 x15, (0 <= x15) ->(0 <= x16) ->(0 <= x17) ->0 <= P_id_match_1 x17 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_0_bounded : forall x16 x17 x15, (0 <= x15) ->(0 <= x16) ->(0 <= x17) ->0 <= P_id_match_0 x17 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_quick_bounded : forall x15, (0 <= x15) ->0 <= P_id_quick x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_False_bounded : 0 <= P_id_False . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_2_bounded : forall x16 x18 x17 x19 x15, (0 <= x15) -> (0 <= x16) -> (0 <= x17) -> (0 <= x18) ->(0 <= x19) ->0 <= P_id_match_2 x19 x18 x17 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Cons_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_Cons x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_5_bounded : forall x16 x18 x17 x15, (0 <= x15) -> (0 <= x16) -> (0 <= x17) ->(0 <= x18) ->0 <= P_id_match_5 x18 x17 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_True_bounded : 0 <= P_id_True . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Pair_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_Pair x16 x15. 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_match_4_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_match_4 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_append x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_3_bounded : forall x16 x20 x18 x17 x21 x19 x15, (0 <= x15) -> (0 <= x16) -> (0 <= x17) -> (0 <= x18) -> (0 <= x19) -> (0 <= x20) -> (0 <= x21) ->0 <= P_id_match_3 x21 x20 x19 x18 x17 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_part_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_part x16 x15. 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_test P_id_match_1 P_id_match_0 P_id_quick P_id_False P_id_match_2 P_id_Cons P_id_match_5 P_id_True P_id_Pair P_id_Nil P_id_match_4 P_id_append P_id_match_3 P_id_part. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_test (x16::x15::nil)) => P_id_test (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_1 (x17::x16:: x15::nil)) => P_id_match_1 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_0 (x17::x16:: x15::nil)) => P_id_match_0 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_quick (x15::nil)) => P_id_quick (measure x15) | (algebra.Alg.Term algebra.F.id_False nil) => P_id_False | (algebra.Alg.Term algebra.F.id_match_2 (x19::x18::x17:: x16::x15::nil)) => P_id_match_2 (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_Cons (x16::x15::nil)) => P_id_Cons (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16:: x15::nil)) => P_id_match_5 (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_True nil) => P_id_True | (algebra.Alg.Term algebra.F.id_Pair (x16::x15::nil)) => P_id_Pair (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_Nil nil) => P_id_Nil | (algebra.Alg.Term algebra.F.id_match_4 (x16::x15::nil)) => P_id_match_4 (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_append (x16::x15::nil)) => P_id_append (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_3 (x21::x20::x19:: x18::x17::x16::x15::nil)) => P_id_match_3 (measure x21) (measure x20) (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_part (x16::x15::nil)) => P_id_part (measure x16) (measure x15) | _ => 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_test_monotonic;assumption. intros ;apply P_id_match_1_monotonic;assumption. intros ;apply P_id_match_0_monotonic;assumption. intros ;apply P_id_quick_monotonic;assumption. intros ;apply P_id_match_2_monotonic;assumption. intros ;apply P_id_Cons_monotonic;assumption. intros ;apply P_id_match_5_monotonic;assumption. intros ;apply P_id_Pair_monotonic;assumption. intros ;apply P_id_match_4_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_match_3_monotonic;assumption. intros ;apply P_id_part_monotonic;assumption. intros ;apply P_id_test_bounded;assumption. intros ;apply P_id_match_1_bounded;assumption. intros ;apply P_id_match_0_bounded;assumption. intros ;apply P_id_quick_bounded;assumption. intros ;apply P_id_False_bounded;assumption. intros ;apply P_id_match_2_bounded;assumption. intros ;apply P_id_Cons_bounded;assumption. intros ;apply P_id_match_5_bounded;assumption. intros ;apply P_id_True_bounded;assumption. intros ;apply P_id_Pair_bounded;assumption. intros ;apply P_id_Nil_bounded;assumption. intros ;apply P_id_match_4_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_match_3_bounded;assumption. intros ;apply P_id_part_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_MATCH_0 (x15:Z) (x16:Z) (x17:Z) := 0. Definition P_id_MATCH_5 (x15:Z) (x16:Z) (x17:Z) (x18:Z) := 0. Definition P_id_MATCH_3 (x15:Z) (x16:Z) (x17:Z) (x18:Z) (x19:Z) (x20: Z) (x21:Z) := 0. Definition P_id_MATCH_1 (x15:Z) (x16:Z) (x17:Z) := 2 + 1* x17. Definition P_id_QUICK (x15:Z) := 0. Definition P_id_PART (x15:Z) (x16:Z) := 3 + 1* x16. Definition P_id_TEST (x15:Z) (x16:Z) := 0. Definition P_id_MATCH_2 (x15:Z) (x16:Z) (x17:Z) (x18:Z) (x19:Z) := 0. Definition P_id_MATCH_4 (x15:Z) (x16:Z) := 0. Definition P_id_APPEND (x15:Z) (x16:Z) := 0. Lemma P_id_MATCH_0_monotonic : forall x16 x20 x18 x17 x19 x15, (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_0 x16 x18 x20 <= P_id_MATCH_0 x15 x17 x19. Proof. intros x20 x19 x18 x17 x16 x15. 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_MATCH_5_monotonic : forall x16 x20 x18 x22 x17 x21 x19 x15, (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_5 x16 x18 x20 x22 <= P_id_MATCH_5 x15 x17 x19 x21. Proof. intros x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MATCH_3_monotonic : forall x16 x24 x20 x28 x18 x26 x22 x17 x25 x21 x19 x27 x23 x15, (0 <= x28)/\ (x28 <= x27) -> (0 <= x26)/\ (x26 <= x25) -> (0 <= x24)/\ (x24 <= x23) -> (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_3 x16 x18 x20 x22 x24 x26 x28 <= P_id_MATCH_3 x15 x17 x19 x21 x23 x25 x27. Proof. intros x28 x27 x26 x25 x24 x23 x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. intros [H_13 H_12]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MATCH_1_monotonic : forall x16 x20 x18 x17 x19 x15, (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_1 x16 x18 x20 <= P_id_MATCH_1 x15 x17 x19. Proof. intros x20 x19 x18 x17 x16 x15. 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_QUICK_monotonic : forall x16 x15, (0 <= x16)/\ (x16 <= x15) ->P_id_QUICK x16 <= P_id_QUICK x15. Proof. intros x16 x15. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PART_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_PART x16 x18 <= P_id_PART x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_TEST_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_TEST x16 x18 <= P_id_TEST x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MATCH_2_monotonic : forall x16 x24 x20 x18 x22 x17 x21 x19 x23 x15, (0 <= x24)/\ (x24 <= x23) -> (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_2 x16 x18 x20 x22 x24 <= P_id_MATCH_2 x15 x17 x19 x21 x23. Proof. intros x24 x23 x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MATCH_4_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_4 x16 x18 <= P_id_MATCH_4 x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_APPEND x16 x18 <= P_id_APPEND x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_test P_id_match_1 P_id_match_0 P_id_quick P_id_False P_id_match_2 P_id_Cons P_id_match_5 P_id_True P_id_Pair P_id_Nil P_id_match_4 P_id_append P_id_match_3 P_id_part P_id_MATCH_0 P_id_MATCH_5 P_id_MATCH_3 P_id_MATCH_1 P_id_QUICK P_id_PART P_id_TEST P_id_MATCH_2 P_id_MATCH_4 P_id_APPEND. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_match_0 (x17:: x16::x15::nil)) => P_id_MATCH_0 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_5 (x18:: x17::x16::x15::nil)) => P_id_MATCH_5 (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_3 (x21:: x20::x19::x18::x17::x16::x15::nil)) => P_id_MATCH_3 (measure x21) (measure x20) (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_1 (x17:: x16::x15::nil)) => P_id_MATCH_1 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_quick (x15::nil)) => P_id_QUICK (measure x15) | (algebra.Alg.Term algebra.F.id_part (x16:: x15::nil)) => P_id_PART (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_test (x16:: x15::nil)) => P_id_TEST (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_2 (x19:: x18::x17::x16::x15::nil)) => P_id_MATCH_2 (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_4 (x16:: x15::nil)) => P_id_MATCH_4 (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_append (x16:: x15::nil)) => P_id_APPEND (measure x16) (measure x15) | _ => 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_test_monotonic;assumption. intros ;apply P_id_match_1_monotonic;assumption. intros ;apply P_id_match_0_monotonic;assumption. intros ;apply P_id_quick_monotonic;assumption. intros ;apply P_id_match_2_monotonic;assumption. intros ;apply P_id_Cons_monotonic;assumption. intros ;apply P_id_match_5_monotonic;assumption. intros ;apply P_id_Pair_monotonic;assumption. intros ;apply P_id_match_4_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_match_3_monotonic;assumption. intros ;apply P_id_part_monotonic;assumption. intros ;apply P_id_test_bounded;assumption. intros ;apply P_id_match_1_bounded;assumption. intros ;apply P_id_match_0_bounded;assumption. intros ;apply P_id_quick_bounded;assumption. intros ;apply P_id_False_bounded;assumption. intros ;apply P_id_match_2_bounded;assumption. intros ;apply P_id_Cons_bounded;assumption. intros ;apply P_id_match_5_bounded;assumption. intros ;apply P_id_True_bounded;assumption. intros ;apply P_id_Pair_bounded;assumption. intros ;apply P_id_Nil_bounded;assumption. intros ;apply P_id_match_4_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_match_3_bounded;assumption. intros ;apply P_id_part_bounded;assumption. apply rules_monotonic. intros ;apply P_id_MATCH_0_monotonic;assumption. intros ;apply P_id_MATCH_5_monotonic;assumption. intros ;apply P_id_MATCH_3_monotonic;assumption. intros ;apply P_id_MATCH_1_monotonic;assumption. intros ;apply P_id_QUICK_monotonic;assumption. intros ;apply P_id_PART_monotonic;assumption. intros ;apply P_id_TEST_monotonic;assumption. intros ;apply P_id_MATCH_2_monotonic;assumption. intros ;apply P_id_MATCH_4_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_4. Proof. intros x. apply well_founded_ind with (R:=fun x y => (Zwf.Zwf 0) (marked_measure x) (marked_measure y)). apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. clear x. intros x IHx. repeat ( constructor;inversion 1;subst; full_prove_ineq algebra.Alg.Term ltac:(algebra.Alg_ext.find_replacement ) algebra.EQT_ext.one_step_list_refl_trans_clos marked_measure marked_measure_star_monotonic (Zwf.Zwf 0) (interp.o_Z 0) ltac:(fun _ => R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 ) ltac:(fun _ => rewrite_and_unfold ) ltac:(fun _ => generate_pos_hyp ) ltac:(fun _ => cbv beta iota zeta delta - [Zplus Zmult Zle Zlt] in * ; try (constructor)) IHx ). Qed. End WF_DP_R_xml_0_scc_4. Definition wf_DP_R_xml_0_scc_4 := WF_DP_R_xml_0_scc_4.wf. Lemma acc_DP_R_xml_0_scc_4 : forall x y, (DP_R_xml_0_scc_4 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_4). intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((eapply acc_DP_R_xml_0_non_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))). apply wf_DP_R_xml_0_scc_4. 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 x16 x12 x9 x13 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Cons (x13::x9::nil)) x15) -> DP_R_xml_0_non_scc_5 (algebra.Alg.Term algebra.F.id_part (x13:: x9::nil)) (algebra.Alg.Term algebra.F.id_match_4 (x16::x15::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; (eapply acc_DP_R_xml_0_scc_4; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Inductive DP_R_xml_0_scc_6 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_6_0 : forall x16 x4 x6 x17 x5 x3 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Cons (x5::x6::nil)) x15) -> DP_R_xml_0_scc_6 (algebra.Alg.Term algebra.F.id_append (x6:: x4::nil)) (algebra.Alg.Term algebra.F.id_match_0 (x17::x16::x15::nil)) (* *) | DP_R_xml_0_scc_6_1 : forall x16 x4 x3 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x3 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x4 x15) -> DP_R_xml_0_scc_6 (algebra.Alg.Term algebra.F.id_match_0 (x3::x4:: x3::nil)) (algebra.Alg.Term algebra.F.id_append (x16::x15::nil)) . Module WF_DP_R_xml_0_scc_6. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_test (x15:Z) (x16:Z) := 1. Definition P_id_match_1 (x15:Z) (x16:Z) (x17:Z) := 1 + 1* x17. Definition P_id_match_0 (x15:Z) (x16:Z) (x17:Z) := 1* x16 + 1* x17. Definition P_id_quick (x15:Z) := 1* x15. Definition P_id_False := 1. Definition P_id_match_2 (x15:Z) (x16:Z) (x17:Z) (x18:Z) (x19:Z) := 2 + 1* x19. Definition P_id_Cons (x15:Z) (x16:Z) := 2 + 1* x16. Definition P_id_match_5 (x15:Z) (x16:Z) (x17:Z) (x18:Z) := 1 + 1* x18. Definition P_id_True := 1. Definition P_id_Pair (x15:Z) (x16:Z) := 1 + 1* x15 + 1* x16. Definition P_id_Nil := 0. Definition P_id_match_4 (x15:Z) (x16:Z) := 1* x16. Definition P_id_append (x15:Z) (x16:Z) := 1* x15 + 1* x16. Definition P_id_match_3 (x15:Z) (x16:Z) (x17:Z) (x18:Z) (x19:Z) (x20: Z) (x21:Z) := 1 + 1* x15 + 1* x16 + 2* x21. Definition P_id_part (x15:Z) (x16:Z) := 1 + 1* x16. Lemma P_id_test_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_test x16 x18 <= P_id_test x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_1_monotonic : forall x16 x20 x18 x17 x19 x15, (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_1 x16 x18 x20 <= P_id_match_1 x15 x17 x19. Proof. intros x20 x19 x18 x17 x16 x15. 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_match_0_monotonic : forall x16 x20 x18 x17 x19 x15, (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_0 x16 x18 x20 <= P_id_match_0 x15 x17 x19. Proof. intros x20 x19 x18 x17 x16 x15. 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_quick_monotonic : forall x16 x15, (0 <= x16)/\ (x16 <= x15) ->P_id_quick x16 <= P_id_quick x15. Proof. intros x16 x15. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_2_monotonic : forall x16 x24 x20 x18 x22 x17 x21 x19 x23 x15, (0 <= x24)/\ (x24 <= x23) -> (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_2 x16 x18 x20 x22 x24 <= P_id_match_2 x15 x17 x19 x21 x23. Proof. intros x24 x23 x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Cons_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_Cons x16 x18 <= P_id_Cons x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_5_monotonic : forall x16 x20 x18 x22 x17 x21 x19 x15, (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_5 x16 x18 x20 x22 <= P_id_match_5 x15 x17 x19 x21. Proof. intros x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Pair_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_Pair x16 x18 <= P_id_Pair x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_4_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_4 x16 x18 <= P_id_match_4 x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_append x16 x18 <= P_id_append x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_3_monotonic : forall x16 x24 x20 x28 x18 x26 x22 x17 x25 x21 x19 x27 x23 x15, (0 <= x28)/\ (x28 <= x27) -> (0 <= x26)/\ (x26 <= x25) -> (0 <= x24)/\ (x24 <= x23) -> (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_3 x16 x18 x20 x22 x24 x26 x28 <= P_id_match_3 x15 x17 x19 x21 x23 x25 x27. Proof. intros x28 x27 x26 x25 x24 x23 x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. intros [H_13 H_12]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_part_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_part x16 x18 <= P_id_part x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_test_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_test x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_1_bounded : forall x16 x17 x15, (0 <= x15) ->(0 <= x16) ->(0 <= x17) ->0 <= P_id_match_1 x17 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_0_bounded : forall x16 x17 x15, (0 <= x15) ->(0 <= x16) ->(0 <= x17) ->0 <= P_id_match_0 x17 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_quick_bounded : forall x15, (0 <= x15) ->0 <= P_id_quick x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_False_bounded : 0 <= P_id_False . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_2_bounded : forall x16 x18 x17 x19 x15, (0 <= x15) -> (0 <= x16) -> (0 <= x17) -> (0 <= x18) ->(0 <= x19) ->0 <= P_id_match_2 x19 x18 x17 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Cons_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_Cons x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_5_bounded : forall x16 x18 x17 x15, (0 <= x15) -> (0 <= x16) -> (0 <= x17) ->(0 <= x18) ->0 <= P_id_match_5 x18 x17 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_True_bounded : 0 <= P_id_True . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Pair_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_Pair x16 x15. 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_match_4_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_match_4 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_append x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_3_bounded : forall x16 x20 x18 x17 x21 x19 x15, (0 <= x15) -> (0 <= x16) -> (0 <= x17) -> (0 <= x18) -> (0 <= x19) -> (0 <= x20) -> (0 <= x21) ->0 <= P_id_match_3 x21 x20 x19 x18 x17 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_part_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_part x16 x15. 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_test P_id_match_1 P_id_match_0 P_id_quick P_id_False P_id_match_2 P_id_Cons P_id_match_5 P_id_True P_id_Pair P_id_Nil P_id_match_4 P_id_append P_id_match_3 P_id_part. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_test (x16::x15::nil)) => P_id_test (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_1 (x17::x16:: x15::nil)) => P_id_match_1 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_0 (x17::x16:: x15::nil)) => P_id_match_0 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_quick (x15::nil)) => P_id_quick (measure x15) | (algebra.Alg.Term algebra.F.id_False nil) => P_id_False | (algebra.Alg.Term algebra.F.id_match_2 (x19::x18::x17:: x16::x15::nil)) => P_id_match_2 (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_Cons (x16::x15::nil)) => P_id_Cons (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16:: x15::nil)) => P_id_match_5 (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_True nil) => P_id_True | (algebra.Alg.Term algebra.F.id_Pair (x16::x15::nil)) => P_id_Pair (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_Nil nil) => P_id_Nil | (algebra.Alg.Term algebra.F.id_match_4 (x16::x15::nil)) => P_id_match_4 (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_append (x16::x15::nil)) => P_id_append (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_3 (x21::x20::x19:: x18::x17::x16::x15::nil)) => P_id_match_3 (measure x21) (measure x20) (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_part (x16::x15::nil)) => P_id_part (measure x16) (measure x15) | _ => 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_test_monotonic;assumption. intros ;apply P_id_match_1_monotonic;assumption. intros ;apply P_id_match_0_monotonic;assumption. intros ;apply P_id_quick_monotonic;assumption. intros ;apply P_id_match_2_monotonic;assumption. intros ;apply P_id_Cons_monotonic;assumption. intros ;apply P_id_match_5_monotonic;assumption. intros ;apply P_id_Pair_monotonic;assumption. intros ;apply P_id_match_4_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_match_3_monotonic;assumption. intros ;apply P_id_part_monotonic;assumption. intros ;apply P_id_test_bounded;assumption. intros ;apply P_id_match_1_bounded;assumption. intros ;apply P_id_match_0_bounded;assumption. intros ;apply P_id_quick_bounded;assumption. intros ;apply P_id_False_bounded;assumption. intros ;apply P_id_match_2_bounded;assumption. intros ;apply P_id_Cons_bounded;assumption. intros ;apply P_id_match_5_bounded;assumption. intros ;apply P_id_True_bounded;assumption. intros ;apply P_id_Pair_bounded;assumption. intros ;apply P_id_Nil_bounded;assumption. intros ;apply P_id_match_4_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_match_3_bounded;assumption. intros ;apply P_id_part_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_MATCH_0 (x15:Z) (x16:Z) (x17:Z) := 2 + 1* x17. Definition P_id_MATCH_5 (x15:Z) (x16:Z) (x17:Z) (x18:Z) := 0. Definition P_id_MATCH_3 (x15:Z) (x16:Z) (x17:Z) (x18:Z) (x19:Z) (x20: Z) (x21:Z) := 0. Definition P_id_MATCH_1 (x15:Z) (x16:Z) (x17:Z) := 0. Definition P_id_QUICK (x15:Z) := 0. Definition P_id_PART (x15:Z) (x16:Z) := 0. Definition P_id_TEST (x15:Z) (x16:Z) := 0. Definition P_id_MATCH_2 (x15:Z) (x16:Z) (x17:Z) (x18:Z) (x19:Z) := 0. Definition P_id_MATCH_4 (x15:Z) (x16:Z) := 0. Definition P_id_APPEND (x15:Z) (x16:Z) := 3 + 1* x15. Lemma P_id_MATCH_0_monotonic : forall x16 x20 x18 x17 x19 x15, (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_0 x16 x18 x20 <= P_id_MATCH_0 x15 x17 x19. Proof. intros x20 x19 x18 x17 x16 x15. 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_MATCH_5_monotonic : forall x16 x20 x18 x22 x17 x21 x19 x15, (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_5 x16 x18 x20 x22 <= P_id_MATCH_5 x15 x17 x19 x21. Proof. intros x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MATCH_3_monotonic : forall x16 x24 x20 x28 x18 x26 x22 x17 x25 x21 x19 x27 x23 x15, (0 <= x28)/\ (x28 <= x27) -> (0 <= x26)/\ (x26 <= x25) -> (0 <= x24)/\ (x24 <= x23) -> (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_3 x16 x18 x20 x22 x24 x26 x28 <= P_id_MATCH_3 x15 x17 x19 x21 x23 x25 x27. Proof. intros x28 x27 x26 x25 x24 x23 x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. intros [H_13 H_12]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MATCH_1_monotonic : forall x16 x20 x18 x17 x19 x15, (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_1 x16 x18 x20 <= P_id_MATCH_1 x15 x17 x19. Proof. intros x20 x19 x18 x17 x16 x15. 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_QUICK_monotonic : forall x16 x15, (0 <= x16)/\ (x16 <= x15) ->P_id_QUICK x16 <= P_id_QUICK x15. Proof. intros x16 x15. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PART_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_PART x16 x18 <= P_id_PART x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_TEST_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_TEST x16 x18 <= P_id_TEST x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MATCH_2_monotonic : forall x16 x24 x20 x18 x22 x17 x21 x19 x23 x15, (0 <= x24)/\ (x24 <= x23) -> (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_2 x16 x18 x20 x22 x24 <= P_id_MATCH_2 x15 x17 x19 x21 x23. Proof. intros x24 x23 x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MATCH_4_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_4 x16 x18 <= P_id_MATCH_4 x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_APPEND x16 x18 <= P_id_APPEND x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_test P_id_match_1 P_id_match_0 P_id_quick P_id_False P_id_match_2 P_id_Cons P_id_match_5 P_id_True P_id_Pair P_id_Nil P_id_match_4 P_id_append P_id_match_3 P_id_part P_id_MATCH_0 P_id_MATCH_5 P_id_MATCH_3 P_id_MATCH_1 P_id_QUICK P_id_PART P_id_TEST P_id_MATCH_2 P_id_MATCH_4 P_id_APPEND. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_match_0 (x17:: x16::x15::nil)) => P_id_MATCH_0 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_5 (x18:: x17::x16::x15::nil)) => P_id_MATCH_5 (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_3 (x21:: x20::x19::x18::x17::x16::x15::nil)) => P_id_MATCH_3 (measure x21) (measure x20) (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_1 (x17:: x16::x15::nil)) => P_id_MATCH_1 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_quick (x15::nil)) => P_id_QUICK (measure x15) | (algebra.Alg.Term algebra.F.id_part (x16:: x15::nil)) => P_id_PART (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_test (x16:: x15::nil)) => P_id_TEST (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_2 (x19:: x18::x17::x16::x15::nil)) => P_id_MATCH_2 (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_4 (x16:: x15::nil)) => P_id_MATCH_4 (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_append (x16:: x15::nil)) => P_id_APPEND (measure x16) (measure x15) | _ => 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_test_monotonic;assumption. intros ;apply P_id_match_1_monotonic;assumption. intros ;apply P_id_match_0_monotonic;assumption. intros ;apply P_id_quick_monotonic;assumption. intros ;apply P_id_match_2_monotonic;assumption. intros ;apply P_id_Cons_monotonic;assumption. intros ;apply P_id_match_5_monotonic;assumption. intros ;apply P_id_Pair_monotonic;assumption. intros ;apply P_id_match_4_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_match_3_monotonic;assumption. intros ;apply P_id_part_monotonic;assumption. intros ;apply P_id_test_bounded;assumption. intros ;apply P_id_match_1_bounded;assumption. intros ;apply P_id_match_0_bounded;assumption. intros ;apply P_id_quick_bounded;assumption. intros ;apply P_id_False_bounded;assumption. intros ;apply P_id_match_2_bounded;assumption. intros ;apply P_id_Cons_bounded;assumption. intros ;apply P_id_match_5_bounded;assumption. intros ;apply P_id_True_bounded;assumption. intros ;apply P_id_Pair_bounded;assumption. intros ;apply P_id_Nil_bounded;assumption. intros ;apply P_id_match_4_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_match_3_bounded;assumption. intros ;apply P_id_part_bounded;assumption. apply rules_monotonic. intros ;apply P_id_MATCH_0_monotonic;assumption. intros ;apply P_id_MATCH_5_monotonic;assumption. intros ;apply P_id_MATCH_3_monotonic;assumption. intros ;apply P_id_MATCH_1_monotonic;assumption. intros ;apply P_id_QUICK_monotonic;assumption. intros ;apply P_id_PART_monotonic;assumption. intros ;apply P_id_TEST_monotonic;assumption. intros ;apply P_id_MATCH_2_monotonic;assumption. intros ;apply P_id_MATCH_4_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_6. Proof. intros x. apply well_founded_ind with (R:=fun x y => (Zwf.Zwf 0) (marked_measure x) (marked_measure y)). apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. clear x. intros x IHx. repeat ( constructor;inversion 1;subst; full_prove_ineq algebra.Alg.Term ltac:(algebra.Alg_ext.find_replacement ) algebra.EQT_ext.one_step_list_refl_trans_clos marked_measure marked_measure_star_monotonic (Zwf.Zwf 0) (interp.o_Z 0) ltac:(fun _ => R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 ) ltac:(fun _ => rewrite_and_unfold ) ltac:(fun _ => generate_pos_hyp ) ltac:(fun _ => cbv beta iota zeta delta - [Zplus Zmult Zle Zlt] in * ; try (constructor)) IHx ). Qed. End WF_DP_R_xml_0_scc_6. Definition wf_DP_R_xml_0_scc_6 := WF_DP_R_xml_0_scc_6.wf. Lemma acc_DP_R_xml_0_scc_6 : forall x y, (DP_R_xml_0_scc_6 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_6). intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). apply wf_DP_R_xml_0_scc_6. Qed. Inductive DP_R_xml_0_non_scc_7 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_7_0 : forall x16 x12 x18 x10 x17 x9 x13 x11 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x13 x18) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x9 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Pair (x10::x11::nil)) x15) -> DP_R_xml_0_non_scc_7 (algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Term algebra.F.id_quick (x10::nil))::(algebra.Alg.Term algebra.F.id_Cons (x13::(algebra.Alg.Term algebra.F.id_quick (x11::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16::x15::nil)) . Lemma acc_DP_R_xml_0_non_scc_7 : forall x y, (DP_R_xml_0_non_scc_7 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_6; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Inductive DP_R_xml_0_scc_8 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_8_0 : forall x16 x12 x18 x10 x17 x9 x13 x11 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x13 x18) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x9 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Pair (x10::x11::nil)) x15) -> DP_R_xml_0_scc_8 (algebra.Alg.Term algebra.F.id_quick (x10::nil)) (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16::x15::nil)) (* *) | DP_R_xml_0_scc_8_1 : forall x12 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x15) -> DP_R_xml_0_scc_8 (algebra.Alg.Term algebra.F.id_match_4 (x12:: x12::nil)) (algebra.Alg.Term algebra.F.id_quick (x15::nil)) (* *) | DP_R_xml_0_scc_8_2 : forall x16 x12 x9 x13 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Cons (x13::x9::nil)) x15) -> DP_R_xml_0_scc_8 (algebra.Alg.Term algebra.F.id_match_5 (x13::x9:: x12::(algebra.Alg.Term algebra.F.id_part (x13:: x9::nil))::nil)) (algebra.Alg.Term algebra.F.id_match_4 (x16::x15::nil)) (* *) | DP_R_xml_0_scc_8_3 : forall x16 x12 x18 x10 x17 x9 x13 x11 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x13 x18) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x9 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Pair (x10::x11::nil)) x15) -> DP_R_xml_0_scc_8 (algebra.Alg.Term algebra.F.id_quick (x11::nil)) (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16::x15::nil)) . Module WF_DP_R_xml_0_scc_8. Inductive DP_R_xml_0_scc_8_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_8_large_0 : forall x16 x12 x18 x10 x17 x9 x13 x11 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x13 x18) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x9 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Pair (x10::x11::nil)) x15) -> DP_R_xml_0_scc_8_large (algebra.Alg.Term algebra.F.id_quick (x10::nil)) (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16::x15::nil)) (* *) | DP_R_xml_0_scc_8_large_1 : forall x12 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x15) -> DP_R_xml_0_scc_8_large (algebra.Alg.Term algebra.F.id_match_4 (x12:: x12::nil)) (algebra.Alg.Term algebra.F.id_quick (x15::nil)) (* *) | DP_R_xml_0_scc_8_large_2 : forall x16 x12 x18 x10 x17 x9 x13 x11 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x13 x18) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x9 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Pair (x10::x11::nil)) x15) -> DP_R_xml_0_scc_8_large (algebra.Alg.Term algebra.F.id_quick (x11::nil)) (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16::x15::nil)) . Inductive DP_R_xml_0_scc_8_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_8_strict_0 : forall x16 x12 x9 x13 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Cons (x13::x9::nil)) x15) -> DP_R_xml_0_scc_8_strict (algebra.Alg.Term algebra.F.id_match_5 (x13::x9::x12::(algebra.Alg.Term algebra.F.id_part (x13::x9::nil))::nil)) (algebra.Alg.Term algebra.F.id_match_4 (x16::x15::nil)) . Module WF_DP_R_xml_0_scc_8_large. Inductive DP_R_xml_0_scc_8_large_non_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_8_large_non_scc_1_0 : forall x12 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x15) -> DP_R_xml_0_scc_8_large_non_scc_1 (algebra.Alg.Term algebra.F.id_match_4 (x12:: x12::nil)) (algebra.Alg.Term algebra.F.id_quick (x15::nil)) . Lemma acc_DP_R_xml_0_scc_8_large_non_scc_1 : forall x y, (DP_R_xml_0_scc_8_large_non_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_8.DP_R_xml_0_scc_8_large x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_scc_8_large_non_scc_2 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_8_large_non_scc_2_0 : forall x16 x12 x18 x10 x17 x9 x13 x11 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x13 x18) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x9 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Pair (x10::x11::nil)) x15) -> DP_R_xml_0_scc_8_large_non_scc_2 (algebra.Alg.Term algebra.F.id_quick (x11::nil)) (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16:: x15::nil)) . Lemma acc_DP_R_xml_0_scc_8_large_non_scc_2 : forall x y, (DP_R_xml_0_scc_8_large_non_scc_2 x y) -> Acc WF_DP_R_xml_0_scc_8.DP_R_xml_0_scc_8_large x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_8_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Inductive DP_R_xml_0_scc_8_large_non_scc_3 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_8_large_non_scc_3_0 : forall x16 x12 x18 x10 x17 x9 x13 x11 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x13 x18) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x9 x17) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x12 x16) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_Pair (x10::x11::nil)) x15) -> DP_R_xml_0_scc_8_large_non_scc_3 (algebra.Alg.Term algebra.F.id_quick (x10::nil)) (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16:: x15::nil)) . Lemma acc_DP_R_xml_0_scc_8_large_non_scc_3 : forall x y, (DP_R_xml_0_scc_8_large_non_scc_3 x y) -> Acc WF_DP_R_xml_0_scc_8.DP_R_xml_0_scc_8_large x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_8_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Lemma wf : well_founded WF_DP_R_xml_0_scc_8.DP_R_xml_0_scc_8_large. Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_8_large_non_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_8_large_non_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_8_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_8_large_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||(fail))))). Qed. End WF_DP_R_xml_0_scc_8_large. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_test (x15:Z) (x16:Z) := 0. Definition P_id_match_1 (x15:Z) (x16:Z) (x17:Z) := 1* x17. Definition P_id_match_0 (x15:Z) (x16:Z) (x17:Z) := 1* x16 + 1* x17. Definition P_id_quick (x15:Z) := 1* x15. Definition P_id_False := 0. Definition P_id_match_2 (x15:Z) (x16:Z) (x17:Z) (x18:Z) (x19:Z) := 1 + 1* x19. Definition P_id_Cons (x15:Z) (x16:Z) := 1 + 1* x16. Definition P_id_match_5 (x15:Z) (x16:Z) (x17:Z) (x18:Z) := 1 + 1* x18. Definition P_id_True := 0. Definition P_id_Pair (x15:Z) (x16:Z) := 1* x15 + 1* x16. Definition P_id_Nil := 0. Definition P_id_match_4 (x15:Z) (x16:Z) := 1* x16. Definition P_id_append (x15:Z) (x16:Z) := 1* x15 + 1* x16. Definition P_id_match_3 (x15:Z) (x16:Z) (x17:Z) (x18:Z) (x19:Z) (x20: Z) (x21:Z) := 1 + 1* x15 + 1* x16. Definition P_id_part (x15:Z) (x16:Z) := 1* x16. Lemma P_id_test_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_test x16 x18 <= P_id_test x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_1_monotonic : forall x16 x20 x18 x17 x19 x15, (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_1 x16 x18 x20 <= P_id_match_1 x15 x17 x19. Proof. intros x20 x19 x18 x17 x16 x15. 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_match_0_monotonic : forall x16 x20 x18 x17 x19 x15, (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_0 x16 x18 x20 <= P_id_match_0 x15 x17 x19. Proof. intros x20 x19 x18 x17 x16 x15. 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_quick_monotonic : forall x16 x15, (0 <= x16)/\ (x16 <= x15) ->P_id_quick x16 <= P_id_quick x15. Proof. intros x16 x15. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_2_monotonic : forall x16 x24 x20 x18 x22 x17 x21 x19 x23 x15, (0 <= x24)/\ (x24 <= x23) -> (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_2 x16 x18 x20 x22 x24 <= P_id_match_2 x15 x17 x19 x21 x23. Proof. intros x24 x23 x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Cons_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_Cons x16 x18 <= P_id_Cons x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_5_monotonic : forall x16 x20 x18 x22 x17 x21 x19 x15, (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_5 x16 x18 x20 x22 <= P_id_match_5 x15 x17 x19 x21. Proof. intros x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Pair_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_Pair x16 x18 <= P_id_Pair x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_4_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_4 x16 x18 <= P_id_match_4 x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_append x16 x18 <= P_id_append x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_3_monotonic : forall x16 x24 x20 x28 x18 x26 x22 x17 x25 x21 x19 x27 x23 x15, (0 <= x28)/\ (x28 <= x27) -> (0 <= x26)/\ (x26 <= x25) -> (0 <= x24)/\ (x24 <= x23) -> (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_match_3 x16 x18 x20 x22 x24 x26 x28 <= P_id_match_3 x15 x17 x19 x21 x23 x25 x27. Proof. intros x28 x27 x26 x25 x24 x23 x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. intros [H_13 H_12]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_part_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_part x16 x18 <= P_id_part x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_test_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_test x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_1_bounded : forall x16 x17 x15, (0 <= x15) ->(0 <= x16) ->(0 <= x17) ->0 <= P_id_match_1 x17 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_0_bounded : forall x16 x17 x15, (0 <= x15) ->(0 <= x16) ->(0 <= x17) ->0 <= P_id_match_0 x17 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_quick_bounded : forall x15, (0 <= x15) ->0 <= P_id_quick x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_False_bounded : 0 <= P_id_False . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_2_bounded : forall x16 x18 x17 x19 x15, (0 <= x15) -> (0 <= x16) -> (0 <= x17) -> (0 <= x18) ->(0 <= x19) ->0 <= P_id_match_2 x19 x18 x17 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Cons_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_Cons x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_5_bounded : forall x16 x18 x17 x15, (0 <= x15) -> (0 <= x16) -> (0 <= x17) ->(0 <= x18) ->0 <= P_id_match_5 x18 x17 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_True_bounded : 0 <= P_id_True . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Pair_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_Pair x16 x15. 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_match_4_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_match_4 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_append x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_match_3_bounded : forall x16 x20 x18 x17 x21 x19 x15, (0 <= x15) -> (0 <= x16) -> (0 <= x17) -> (0 <= x18) -> (0 <= x19) -> (0 <= x20) -> (0 <= x21) ->0 <= P_id_match_3 x21 x20 x19 x18 x17 x16 x15. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_part_bounded : forall x16 x15, (0 <= x15) ->(0 <= x16) ->0 <= P_id_part x16 x15. 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_test P_id_match_1 P_id_match_0 P_id_quick P_id_False P_id_match_2 P_id_Cons P_id_match_5 P_id_True P_id_Pair P_id_Nil P_id_match_4 P_id_append P_id_match_3 P_id_part. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_test (x16::x15::nil)) => P_id_test (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_1 (x17::x16:: x15::nil)) => P_id_match_1 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_0 (x17::x16:: x15::nil)) => P_id_match_0 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_quick (x15::nil)) => P_id_quick (measure x15) | (algebra.Alg.Term algebra.F.id_False nil) => P_id_False | (algebra.Alg.Term algebra.F.id_match_2 (x19::x18::x17:: x16::x15::nil)) => P_id_match_2 (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_Cons (x16::x15::nil)) => P_id_Cons (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_5 (x18::x17::x16:: x15::nil)) => P_id_match_5 (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_True nil) => P_id_True | (algebra.Alg.Term algebra.F.id_Pair (x16::x15::nil)) => P_id_Pair (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_Nil nil) => P_id_Nil | (algebra.Alg.Term algebra.F.id_match_4 (x16::x15::nil)) => P_id_match_4 (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_append (x16::x15::nil)) => P_id_append (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_3 (x21::x20::x19:: x18::x17::x16::x15::nil)) => P_id_match_3 (measure x21) (measure x20) (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_part (x16::x15::nil)) => P_id_part (measure x16) (measure x15) | _ => 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_test_monotonic;assumption. intros ;apply P_id_match_1_monotonic;assumption. intros ;apply P_id_match_0_monotonic;assumption. intros ;apply P_id_quick_monotonic;assumption. intros ;apply P_id_match_2_monotonic;assumption. intros ;apply P_id_Cons_monotonic;assumption. intros ;apply P_id_match_5_monotonic;assumption. intros ;apply P_id_Pair_monotonic;assumption. intros ;apply P_id_match_4_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_match_3_monotonic;assumption. intros ;apply P_id_part_monotonic;assumption. intros ;apply P_id_test_bounded;assumption. intros ;apply P_id_match_1_bounded;assumption. intros ;apply P_id_match_0_bounded;assumption. intros ;apply P_id_quick_bounded;assumption. intros ;apply P_id_False_bounded;assumption. intros ;apply P_id_match_2_bounded;assumption. intros ;apply P_id_Cons_bounded;assumption. intros ;apply P_id_match_5_bounded;assumption. intros ;apply P_id_True_bounded;assumption. intros ;apply P_id_Pair_bounded;assumption. intros ;apply P_id_Nil_bounded;assumption. intros ;apply P_id_match_4_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_match_3_bounded;assumption. intros ;apply P_id_part_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_MATCH_0 (x15:Z) (x16:Z) (x17:Z) := 0. Definition P_id_MATCH_5 (x15:Z) (x16:Z) (x17:Z) (x18:Z) := 2* x18. Definition P_id_MATCH_3 (x15:Z) (x16:Z) (x17:Z) (x18:Z) (x19:Z) (x20: Z) (x21:Z) := 0. Definition P_id_MATCH_1 (x15:Z) (x16:Z) (x17:Z) := 0. Definition P_id_QUICK (x15:Z) := 2* x15. Definition P_id_PART (x15:Z) (x16:Z) := 0. Definition P_id_TEST (x15:Z) (x16:Z) := 0. Definition P_id_MATCH_2 (x15:Z) (x16:Z) (x17:Z) (x18:Z) (x19:Z) := 0. Definition P_id_MATCH_4 (x15:Z) (x16:Z) := 2* x16. Definition P_id_APPEND (x15:Z) (x16:Z) := 0. Lemma P_id_MATCH_0_monotonic : forall x16 x20 x18 x17 x19 x15, (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_0 x16 x18 x20 <= P_id_MATCH_0 x15 x17 x19. Proof. intros x20 x19 x18 x17 x16 x15. 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_MATCH_5_monotonic : forall x16 x20 x18 x22 x17 x21 x19 x15, (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_5 x16 x18 x20 x22 <= P_id_MATCH_5 x15 x17 x19 x21. Proof. intros x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MATCH_3_monotonic : forall x16 x24 x20 x28 x18 x26 x22 x17 x25 x21 x19 x27 x23 x15, (0 <= x28)/\ (x28 <= x27) -> (0 <= x26)/\ (x26 <= x25) -> (0 <= x24)/\ (x24 <= x23) -> (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_3 x16 x18 x20 x22 x24 x26 x28 <= P_id_MATCH_3 x15 x17 x19 x21 x23 x25 x27. Proof. intros x28 x27 x26 x25 x24 x23 x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. intros [H_13 H_12]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MATCH_1_monotonic : forall x16 x20 x18 x17 x19 x15, (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_1 x16 x18 x20 <= P_id_MATCH_1 x15 x17 x19. Proof. intros x20 x19 x18 x17 x16 x15. 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_QUICK_monotonic : forall x16 x15, (0 <= x16)/\ (x16 <= x15) ->P_id_QUICK x16 <= P_id_QUICK x15. Proof. intros x16 x15. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PART_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_PART x16 x18 <= P_id_PART x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_TEST_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_TEST x16 x18 <= P_id_TEST x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MATCH_2_monotonic : forall x16 x24 x20 x18 x22 x17 x21 x19 x23 x15, (0 <= x24)/\ (x24 <= x23) -> (0 <= x22)/\ (x22 <= x21) -> (0 <= x20)/\ (x20 <= x19) -> (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_2 x16 x18 x20 x22 x24 <= P_id_MATCH_2 x15 x17 x19 x21 x23. Proof. intros x24 x23 x22 x21 x20 x19 x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MATCH_4_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) -> P_id_MATCH_4 x16 x18 <= P_id_MATCH_4 x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x16 x18 x17 x15, (0 <= x18)/\ (x18 <= x17) -> (0 <= x16)/\ (x16 <= x15) ->P_id_APPEND x16 x18 <= P_id_APPEND x15 x17. Proof. intros x18 x17 x16 x15. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_test P_id_match_1 P_id_match_0 P_id_quick P_id_False P_id_match_2 P_id_Cons P_id_match_5 P_id_True P_id_Pair P_id_Nil P_id_match_4 P_id_append P_id_match_3 P_id_part P_id_MATCH_0 P_id_MATCH_5 P_id_MATCH_3 P_id_MATCH_1 P_id_QUICK P_id_PART P_id_TEST P_id_MATCH_2 P_id_MATCH_4 P_id_APPEND. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_match_0 (x17:: x16::x15::nil)) => P_id_MATCH_0 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_5 (x18:: x17::x16::x15::nil)) => P_id_MATCH_5 (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_3 (x21:: x20::x19::x18::x17::x16::x15::nil)) => P_id_MATCH_3 (measure x21) (measure x20) (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_1 (x17:: x16::x15::nil)) => P_id_MATCH_1 (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_quick (x15::nil)) => P_id_QUICK (measure x15) | (algebra.Alg.Term algebra.F.id_part (x16:: x15::nil)) => P_id_PART (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_test (x16:: x15::nil)) => P_id_TEST (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_2 (x19:: x18::x17::x16::x15::nil)) => P_id_MATCH_2 (measure x19) (measure x18) (measure x17) (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_match_4 (x16:: x15::nil)) => P_id_MATCH_4 (measure x16) (measure x15) | (algebra.Alg.Term algebra.F.id_append (x16:: x15::nil)) => P_id_APPEND (measure x16) (measure x15) | _ => 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_test_monotonic;assumption. intros ;apply P_id_match_1_monotonic;assumption. intros ;apply P_id_match_0_monotonic;assumption. intros ;apply P_id_quick_monotonic;assumption. intros ;apply P_id_match_2_monotonic;assumption. intros ;apply P_id_Cons_monotonic;assumption. intros ;apply P_id_match_5_monotonic;assumption. intros ;apply P_id_Pair_monotonic;assumption. intros ;apply P_id_match_4_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_match_3_monotonic;assumption. intros ;apply P_id_part_monotonic;assumption. intros ;apply P_id_test_bounded;assumption. intros ;apply P_id_match_1_bounded;assumption. intros ;apply P_id_match_0_bounded;assumption. intros ;apply P_id_quick_bounded;assumption. intros ;apply P_id_False_bounded;assumption. intros ;apply P_id_match_2_bounded;assumption. intros ;apply P_id_Cons_bounded;assumption. intros ;apply P_id_match_5_bounded;assumption. intros ;apply P_id_True_bounded;assumption. intros ;apply P_id_Pair_bounded;assumption. intros ;apply P_id_Nil_bounded;assumption. intros ;apply P_id_match_4_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_match_3_bounded;assumption. intros ;apply P_id_part_bounded;assumption. apply rules_monotonic. intros ;apply P_id_MATCH_0_monotonic;assumption. intros ;apply P_id_MATCH_5_monotonic;assumption. intros ;apply P_id_MATCH_3_monotonic;assumption. intros ;apply P_id_MATCH_1_monotonic;assumption. intros ;apply P_id_QUICK_monotonic;assumption. intros ;apply P_id_PART_monotonic;assumption. intros ;apply P_id_TEST_monotonic;assumption. intros ;apply P_id_MATCH_2_monotonic;assumption. intros ;apply P_id_MATCH_4_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b). Definition le a b := marked_measure a <= marked_measure b. Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c. Proof. unfold lt, le in *. intros a b c. apply (interp.le_lt_compat_right (interp.o_Z 0)). Qed. Lemma wf_lt : well_founded lt. Proof. unfold lt in *. apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. Qed. Lemma DP_R_xml_0_scc_8_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_8_strict lt. Proof. unfold Relation_Definitions.inclusion, lt in *. intros a b H;destruct H; match goal with | |- (Zwf.Zwf 0) _ (marked_measure (algebra.Alg.Term ?f ?l)) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_lt_compat_right (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma DP_R_xml_0_scc_8_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_8_large le. Proof. unfold Relation_Definitions.inclusion, le, Zwf.Zwf in *. intros a b H;destruct H; match goal with | |- _ <= marked_measure (algebra.Alg.Term ?f ?l) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_trans (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition wf_DP_R_xml_0_scc_8_large := WF_DP_R_xml_0_scc_8_large.wf. Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_8. Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_8_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx';apply DP_R_xml_0_scc_8_strict_in_lt; econstructor eassumption)|| ((apply IHx;[econstructor eassumption| intros y' Hlt;apply IHx';apply lt_le_compat with (1:=Hlt) ; apply DP_R_xml_0_scc_8_large_in_le;econstructor eassumption])). apply wf_DP_R_xml_0_scc_8_large. Qed. End WF_DP_R_xml_0_scc_8. Definition wf_DP_R_xml_0_scc_8 := WF_DP_R_xml_0_scc_8.wf. Lemma acc_DP_R_xml_0_scc_8 : forall x y, (DP_R_xml_0_scc_8 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_8). intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((eapply acc_DP_R_xml_0_non_scc_7; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_5; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). apply wf_DP_R_xml_0_scc_8. 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_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_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___Cime___quick.trs/a3pat.v" *** *** End: *** *)