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_0 *) | id_0 : symb (* id_bs *) | id_bs : symb (* id_and *) | id_and : symb (* id__minus_ *) | id__minus_ : symb (* id_l *) | id_l : symb (* id__plus_ *) | id__plus_ : symb (* id_wb *) | id_wb : symb (* id_ge *) | id_ge : symb (* id_false *) | id_false : symb (* id_min *) | id_min : symb (* id__sharp_ *) | id__sharp_ : symb (* id_size *) | id_size : symb (* id_if *) | id_if : symb (* id_not *) | id_not : symb (* id_n *) | id_n : symb (* id_1 *) | id_1 : symb (* id_val *) | id_val : symb (* id_true *) | id_true : symb (* id_max *) | id_max : symb . Definition symb_eq_bool (f1 f2:symb) : bool := match f1,f2 with | id_0,id_0 => true | id_bs,id_bs => true | id_and,id_and => true | id__minus_,id__minus_ => true | id_l,id_l => true | id__plus_,id__plus_ => true | id_wb,id_wb => true | id_ge,id_ge => true | id_false,id_false => true | id_min,id_min => true | id__sharp_,id__sharp_ => true | id_size,id_size => true | id_if,id_if => true | id_not,id_not => true | id_n,id_n => true | id_1,id_1 => true | id_val,id_val => true | id_true,id_true => true | id_max,id_max => 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_0,id_0 => refl_equal _ | id_bs,id_bs => refl_equal _ | id_and,id_and => refl_equal _ | id__minus_,id__minus_ => refl_equal _ | id_l,id_l => refl_equal _ | id__plus_,id__plus_ => refl_equal _ | id_wb,id_wb => refl_equal _ | id_ge,id_ge => refl_equal _ | id_false,id_false => refl_equal _ | id_min,id_min => refl_equal _ | id__sharp_,id__sharp_ => refl_equal _ | id_size,id_size => refl_equal _ | id_if,id_if => refl_equal _ | id_not,id_not => refl_equal _ | id_n,id_n => refl_equal _ | id_1,id_1 => refl_equal _ | id_val,id_val => refl_equal _ | id_true,id_true => refl_equal _ | id_max,id_max => refl_equal _ | _,_ => _ end;intros abs;discriminate. Defined. Definition arity (f:symb) := match f with | id_0 => term.Free 1 | id_bs => term.Free 1 | id_and => term.Free 2 | id__minus_ => term.Free 2 | id_l => term.Free 1 | id__plus_ => term.Free 2 | id_wb => term.Free 1 | id_ge => term.Free 2 | id_false => term.Free 0 | id_min => term.Free 1 | id__sharp_ => term.Free 0 | id_size => term.Free 1 | id_if => term.Free 3 | id_not => term.Free 1 | id_n => term.Free 3 | id_1 => term.Free 1 | id_val => term.Free 1 | id_true => term.Free 0 | id_max => term.Free 1 end. Definition symb_order (f1 f2:symb) : bool := match f1,f2 with | id_0,id_0 => true | id_0,id_bs => false | id_0,id_and => false | id_0,id__minus_ => false | id_0,id_l => false | id_0,id__plus_ => false | id_0,id_wb => false | id_0,id_ge => false | id_0,id_false => false | id_0,id_min => false | id_0,id__sharp_ => false | id_0,id_size => false | id_0,id_if => false | id_0,id_not => false | id_0,id_n => false | id_0,id_1 => false | id_0,id_val => false | id_0,id_true => false | id_0,id_max => false | id_bs,id_0 => true | id_bs,id_bs => true | id_bs,id_and => false | id_bs,id__minus_ => false | id_bs,id_l => false | id_bs,id__plus_ => false | id_bs,id_wb => false | id_bs,id_ge => false | id_bs,id_false => false | id_bs,id_min => false | id_bs,id__sharp_ => false | id_bs,id_size => false | id_bs,id_if => false | id_bs,id_not => false | id_bs,id_n => false | id_bs,id_1 => false | id_bs,id_val => false | id_bs,id_true => false | id_bs,id_max => false | id_and,id_0 => true | id_and,id_bs => true | id_and,id_and => true | id_and,id__minus_ => false | id_and,id_l => false | id_and,id__plus_ => false | id_and,id_wb => false | id_and,id_ge => false | id_and,id_false => false | id_and,id_min => false | id_and,id__sharp_ => false | id_and,id_size => false | id_and,id_if => false | id_and,id_not => false | id_and,id_n => false | id_and,id_1 => false | id_and,id_val => false | id_and,id_true => false | id_and,id_max => false | id__minus_,id_0 => true | id__minus_,id_bs => true | id__minus_,id_and => true | id__minus_,id__minus_ => true | id__minus_,id_l => false | id__minus_,id__plus_ => false | id__minus_,id_wb => false | id__minus_,id_ge => false | id__minus_,id_false => false | id__minus_,id_min => false | id__minus_,id__sharp_ => false | id__minus_,id_size => false | id__minus_,id_if => false | id__minus_,id_not => false | id__minus_,id_n => false | id__minus_,id_1 => false | id__minus_,id_val => false | id__minus_,id_true => false | id__minus_,id_max => false | id_l,id_0 => true | id_l,id_bs => true | id_l,id_and => true | id_l,id__minus_ => true | id_l,id_l => true | id_l,id__plus_ => false | id_l,id_wb => false | id_l,id_ge => false | id_l,id_false => false | id_l,id_min => false | id_l,id__sharp_ => false | id_l,id_size => false | id_l,id_if => false | id_l,id_not => false | id_l,id_n => false | id_l,id_1 => false | id_l,id_val => false | id_l,id_true => false | id_l,id_max => false | id__plus_,id_0 => true | id__plus_,id_bs => true | id__plus_,id_and => true | id__plus_,id__minus_ => true | id__plus_,id_l => true | id__plus_,id__plus_ => true | id__plus_,id_wb => false | id__plus_,id_ge => false | id__plus_,id_false => false | id__plus_,id_min => false | id__plus_,id__sharp_ => false | id__plus_,id_size => false | id__plus_,id_if => false | id__plus_,id_not => false | id__plus_,id_n => false | id__plus_,id_1 => false | id__plus_,id_val => false | id__plus_,id_true => false | id__plus_,id_max => false | id_wb,id_0 => true | id_wb,id_bs => true | id_wb,id_and => true | id_wb,id__minus_ => true | id_wb,id_l => true | id_wb,id__plus_ => true | id_wb,id_wb => true | id_wb,id_ge => false | id_wb,id_false => false | id_wb,id_min => false | id_wb,id__sharp_ => false | id_wb,id_size => false | id_wb,id_if => false | id_wb,id_not => false | id_wb,id_n => false | id_wb,id_1 => false | id_wb,id_val => false | id_wb,id_true => false | id_wb,id_max => false | id_ge,id_0 => true | id_ge,id_bs => true | id_ge,id_and => true | id_ge,id__minus_ => true | id_ge,id_l => true | id_ge,id__plus_ => true | id_ge,id_wb => true | id_ge,id_ge => true | id_ge,id_false => false | id_ge,id_min => false | id_ge,id__sharp_ => false | id_ge,id_size => false | id_ge,id_if => false | id_ge,id_not => false | id_ge,id_n => false | id_ge,id_1 => false | id_ge,id_val => false | id_ge,id_true => false | id_ge,id_max => false | id_false,id_0 => true | id_false,id_bs => true | id_false,id_and => true | id_false,id__minus_ => true | id_false,id_l => true | id_false,id__plus_ => true | id_false,id_wb => true | id_false,id_ge => true | id_false,id_false => true | id_false,id_min => false | id_false,id__sharp_ => false | id_false,id_size => false | id_false,id_if => false | id_false,id_not => false | id_false,id_n => false | id_false,id_1 => false | id_false,id_val => false | id_false,id_true => false | id_false,id_max => false | id_min,id_0 => true | id_min,id_bs => true | id_min,id_and => true | id_min,id__minus_ => true | id_min,id_l => true | id_min,id__plus_ => true | id_min,id_wb => true | id_min,id_ge => true | id_min,id_false => true | id_min,id_min => true | id_min,id__sharp_ => false | id_min,id_size => false | id_min,id_if => false | id_min,id_not => false | id_min,id_n => false | id_min,id_1 => false | id_min,id_val => false | id_min,id_true => false | id_min,id_max => false | id__sharp_,id_0 => true | id__sharp_,id_bs => true | id__sharp_,id_and => true | id__sharp_,id__minus_ => true | id__sharp_,id_l => true | id__sharp_,id__plus_ => true | id__sharp_,id_wb => true | id__sharp_,id_ge => true | id__sharp_,id_false => true | id__sharp_,id_min => true | id__sharp_,id__sharp_ => true | id__sharp_,id_size => false | id__sharp_,id_if => false | id__sharp_,id_not => false | id__sharp_,id_n => false | id__sharp_,id_1 => false | id__sharp_,id_val => false | id__sharp_,id_true => false | id__sharp_,id_max => false | id_size,id_0 => true | id_size,id_bs => true | id_size,id_and => true | id_size,id__minus_ => true | id_size,id_l => true | id_size,id__plus_ => true | id_size,id_wb => true | id_size,id_ge => true | id_size,id_false => true | id_size,id_min => true | id_size,id__sharp_ => true | id_size,id_size => true | id_size,id_if => false | id_size,id_not => false | id_size,id_n => false | id_size,id_1 => false | id_size,id_val => false | id_size,id_true => false | id_size,id_max => false | id_if,id_0 => true | id_if,id_bs => true | id_if,id_and => true | id_if,id__minus_ => true | id_if,id_l => true | id_if,id__plus_ => true | id_if,id_wb => true | id_if,id_ge => true | id_if,id_false => true | id_if,id_min => true | id_if,id__sharp_ => true | id_if,id_size => true | id_if,id_if => true | id_if,id_not => false | id_if,id_n => false | id_if,id_1 => false | id_if,id_val => false | id_if,id_true => false | id_if,id_max => false | id_not,id_0 => true | id_not,id_bs => true | id_not,id_and => true | id_not,id__minus_ => true | id_not,id_l => true | id_not,id__plus_ => true | id_not,id_wb => true | id_not,id_ge => true | id_not,id_false => true | id_not,id_min => true | id_not,id__sharp_ => true | id_not,id_size => true | id_not,id_if => true | id_not,id_not => true | id_not,id_n => false | id_not,id_1 => false | id_not,id_val => false | id_not,id_true => false | id_not,id_max => false | id_n,id_0 => true | id_n,id_bs => true | id_n,id_and => true | id_n,id__minus_ => true | id_n,id_l => true | id_n,id__plus_ => true | id_n,id_wb => true | id_n,id_ge => true | id_n,id_false => true | id_n,id_min => true | id_n,id__sharp_ => true | id_n,id_size => true | id_n,id_if => true | id_n,id_not => true | id_n,id_n => true | id_n,id_1 => false | id_n,id_val => false | id_n,id_true => false | id_n,id_max => false | id_1,id_0 => true | id_1,id_bs => true | id_1,id_and => true | id_1,id__minus_ => true | id_1,id_l => true | id_1,id__plus_ => true | id_1,id_wb => true | id_1,id_ge => true | id_1,id_false => true | id_1,id_min => true | id_1,id__sharp_ => true | id_1,id_size => true | id_1,id_if => true | id_1,id_not => true | id_1,id_n => true | id_1,id_1 => true | id_1,id_val => false | id_1,id_true => false | id_1,id_max => false | id_val,id_0 => true | id_val,id_bs => true | id_val,id_and => true | id_val,id__minus_ => true | id_val,id_l => true | id_val,id__plus_ => true | id_val,id_wb => true | id_val,id_ge => true | id_val,id_false => true | id_val,id_min => true | id_val,id__sharp_ => true | id_val,id_size => true | id_val,id_if => true | id_val,id_not => true | id_val,id_n => true | id_val,id_1 => true | id_val,id_val => true | id_val,id_true => false | id_val,id_max => false | id_true,id_0 => true | id_true,id_bs => true | id_true,id_and => true | id_true,id__minus_ => true | id_true,id_l => true | id_true,id__plus_ => true | id_true,id_wb => true | id_true,id_ge => true | id_true,id_false => true | id_true,id_min => true | id_true,id__sharp_ => true | id_true,id_size => true | id_true,id_if => true | id_true,id_not => true | id_true,id_n => true | id_true,id_1 => true | id_true,id_val => true | id_true,id_true => true | id_true,id_max => false | id_max,id_0 => true | id_max,id_bs => true | id_max,id_and => true | id_max,id__minus_ => true | id_max,id_l => true | id_max,id__plus_ => true | id_max,id_wb => true | id_max,id_ge => true | id_max,id_false => true | id_max,id_min => true | id_max,id__sharp_ => true | id_max,id_size => true | id_max,id_if => true | id_max,id_not => true | id_max,id_n => true | id_max,id_1 => true | id_max,id_val => true | id_max,id_true => true | id_max,id_max => 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 := (* 0(#) -> # *) | R_xml_0_rule_0 : R_xml_0_rules (algebra.Alg.Term algebra.F.id__sharp_ nil) (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil)) (* +(x_,#) -> x_ *) | R_xml_0_rule_1 : R_xml_0_rules (algebra.Alg.Var 1) (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id__sharp_ nil)::nil)) (* +(#,x_) -> x_ *) | R_xml_0_rule_2 : R_xml_0_rules (algebra.Alg.Var 1) (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__sharp_ nil)::(algebra.Alg.Var 1)::nil)) (* +(0(x_),0(y_)) -> 0(+(x_,y_)) *) | R_xml_0_rule_3 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil))::nil)) (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 2)::nil))::nil)) (* +(0(x_),1(y_)) -> 1(+(x_,y_)) *) | R_xml_0_rule_4 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil))::nil)) (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 2)::nil))::nil)) (* +(1(x_),0(y_)) -> 1(+(x_,y_)) *) | R_xml_0_rule_5 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil))::nil)) (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 2)::nil))::nil)) (* +(1(x_),1(y_)) -> 0(+(+(x_,y_),1(#))) *) | R_xml_0_rule_6 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 2)::nil))::nil)) (* +(x_,+(y_,z_)) -> +(+(x_,y_),z_) *) | R_xml_0_rule_7 : R_xml_0_rules (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil))::(algebra.Alg.Var 3)::nil)) (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Var 2):: (algebra.Alg.Var 3)::nil))::nil)) (* -(x_,#) -> x_ *) | R_xml_0_rule_8 : R_xml_0_rules (algebra.Alg.Var 1) (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id__sharp_ nil)::nil)) (* -(#,x_) -> # *) | R_xml_0_rule_9 : R_xml_0_rules (algebra.Alg.Term algebra.F.id__sharp_ nil) (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id__sharp_ nil)::(algebra.Alg.Var 1)::nil)) (* -(0(x_),0(y_)) -> 0(-(x_,y_)) *) | R_xml_0_rule_10 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil))::nil)) (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 2)::nil))::nil)) (* -(0(x_),1(y_)) -> 1(-(-(x_,y_),1(#))) *) | R_xml_0_rule_11 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 2)::nil))::nil)) (* -(1(x_),0(y_)) -> 1(-(x_,y_)) *) | R_xml_0_rule_12 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil))::nil)) (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 2)::nil))::nil)) (* -(1(x_),1(y_)) -> 0(-(x_,y_)) *) | R_xml_0_rule_13 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil))::nil)) (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 2)::nil))::nil)) (* not(false) -> true *) | R_xml_0_rule_14 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_true nil) (algebra.Alg.Term algebra.F.id_not ((algebra.Alg.Term algebra.F.id_false nil)::nil)) (* not(true) -> false *) | R_xml_0_rule_15 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_false nil) (algebra.Alg.Term algebra.F.id_not ((algebra.Alg.Term algebra.F.id_true nil)::nil)) (* and(x_,true) -> x_ *) | R_xml_0_rule_16 : R_xml_0_rules (algebra.Alg.Var 1) (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id_true nil)::nil)) (* and(x_,false) -> false *) | R_xml_0_rule_17 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_false nil) (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id_false nil)::nil)) (* if(true,x_,y_) -> x_ *) | R_xml_0_rule_18 : R_xml_0_rules (algebra.Alg.Var 1) (algebra.Alg.Term algebra.F.id_if ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil)) (* if(false,x_,y_) -> y_ *) | R_xml_0_rule_19 : R_xml_0_rules (algebra.Alg.Var 2) (algebra.Alg.Term algebra.F.id_if ((algebra.Alg.Term algebra.F.id_false nil)::(algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil)) (* ge(0(x_),0(y_)) -> ge(x_,y_) *) | R_xml_0_rule_20 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil)) (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 2)::nil))::nil)) (* ge(0(x_),1(y_)) -> not(ge(y_,x_)) *) | R_xml_0_rule_21 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_not ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Var 2):: (algebra.Alg.Var 1)::nil))::nil)) (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 2)::nil))::nil)) (* ge(1(x_),0(y_)) -> ge(x_,y_) *) | R_xml_0_rule_22 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil)) (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 2)::nil))::nil)) (* ge(1(x_),1(y_)) -> ge(x_,y_) *) | R_xml_0_rule_23 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil)) (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 2)::nil))::nil)) (* ge(x_,#) -> true *) | R_xml_0_rule_24 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_true nil) (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id__sharp_ nil)::nil)) (* ge(#,1(x_)) -> false *) | R_xml_0_rule_25 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_false nil) (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id__sharp_ nil)::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 1)::nil))::nil)) (* ge(#,0(x_)) -> ge(#,x_) *) | R_xml_0_rule_26 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id__sharp_ nil)::(algebra.Alg.Var 1)::nil)) (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id__sharp_ nil)::(algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 1)::nil))::nil)) (* val(l(x_)) -> x_ *) | R_xml_0_rule_27 : R_xml_0_rules (algebra.Alg.Var 1) (algebra.Alg.Term algebra.F.id_val ((algebra.Alg.Term algebra.F.id_l ((algebra.Alg.Var 1)::nil))::nil)) (* val(n(x_,y_,z_)) -> x_ *) | R_xml_0_rule_28 : R_xml_0_rules (algebra.Alg.Var 1) (algebra.Alg.Term algebra.F.id_val ((algebra.Alg.Term algebra.F.id_n ((algebra.Alg.Var 1)::(algebra.Alg.Var 2):: (algebra.Alg.Var 3)::nil))::nil)) (* min(l(x_)) -> x_ *) | R_xml_0_rule_29 : R_xml_0_rules (algebra.Alg.Var 1) (algebra.Alg.Term algebra.F.id_min ((algebra.Alg.Term algebra.F.id_l ((algebra.Alg.Var 1)::nil))::nil)) (* min(n(x_,y_,z_)) -> min(y_) *) | R_xml_0_rule_30 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_min ((algebra.Alg.Var 2)::nil)) (algebra.Alg.Term algebra.F.id_min ((algebra.Alg.Term algebra.F.id_n ((algebra.Alg.Var 1)::(algebra.Alg.Var 2):: (algebra.Alg.Var 3)::nil))::nil)) (* max(l(x_)) -> x_ *) | R_xml_0_rule_31 : R_xml_0_rules (algebra.Alg.Var 1) (algebra.Alg.Term algebra.F.id_max ((algebra.Alg.Term algebra.F.id_l ((algebra.Alg.Var 1)::nil))::nil)) (* max(n(x_,y_,z_)) -> max(z_) *) | R_xml_0_rule_32 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_max ((algebra.Alg.Var 3)::nil)) (algebra.Alg.Term algebra.F.id_max ((algebra.Alg.Term algebra.F.id_n ((algebra.Alg.Var 1)::(algebra.Alg.Var 2):: (algebra.Alg.Var 3)::nil))::nil)) (* bs(l(x_)) -> true *) | R_xml_0_rule_33 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_true nil) (algebra.Alg.Term algebra.F.id_bs ((algebra.Alg.Term algebra.F.id_l ((algebra.Alg.Var 1)::nil))::nil)) (* bs(n(x_,y_,z_)) -> and(and(ge(x_,max(y_)),ge(min(z_),x_)),and(bs(y_),bs(z_))) *) | R_xml_0_rule_34 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Var 1)::(algebra.Alg.Term algebra.F.id_max ((algebra.Alg.Var 2)::nil))::nil))::(algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_min ((algebra.Alg.Var 3)::nil)):: (algebra.Alg.Var 1)::nil))::nil))::(algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_bs ((algebra.Alg.Var 2)::nil))::(algebra.Alg.Term algebra.F.id_bs ((algebra.Alg.Var 3)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_bs ((algebra.Alg.Term algebra.F.id_n ((algebra.Alg.Var 1)::(algebra.Alg.Var 2):: (algebra.Alg.Var 3)::nil))::nil)) (* size(l(x_)) -> 1(#) *) | R_xml_0_rule_35 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil)) (algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Term algebra.F.id_l ((algebra.Alg.Var 1)::nil))::nil)) (* size(n(x_,y_,z_)) -> +(+(size(x_),size(y_)),1(#)) *) | R_xml_0_rule_36 : R_xml_0_rules (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Var 2)::nil))::nil)):: (algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Term algebra.F.id_n ((algebra.Alg.Var 1)::(algebra.Alg.Var 2):: (algebra.Alg.Var 3)::nil))::nil)) (* wb(l(x_)) -> true *) | R_xml_0_rule_37 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_true nil) (algebra.Alg.Term algebra.F.id_wb ((algebra.Alg.Term algebra.F.id_l ((algebra.Alg.Var 1)::nil))::nil)) (* wb(n(x_,y_,z_)) -> and(if(ge(size(y_),size(z_)),ge(1(#),-(size(y_),size(z_))), ge(1(#),-(size(z_),size(y_)))),and(wb(y_),wb(z_))) *) | R_xml_0_rule_38 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_if ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Var 2)::nil))::(algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Var 3)::nil))::nil)):: (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::(algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Var 2)::nil))::(algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Var 3)::nil))::nil))::nil)):: (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::(algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Var 3)::nil))::(algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Var 2)::nil))::nil))::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_wb ((algebra.Alg.Var 2)::nil)):: (algebra.Alg.Term algebra.F.id_wb ((algebra.Alg.Var 3)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb ((algebra.Alg.Term algebra.F.id_n ((algebra.Alg.Var 1)::(algebra.Alg.Var 2):: (algebra.Alg.Var 3)::nil))::nil)) . Definition R_xml_0_rule_as_list_0 := ((algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil)),(algebra.Alg.Term algebra.F.id__sharp_ nil))::nil. Definition R_xml_0_rule_as_list_1 := ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id__sharp_ nil)::nil)),(algebra.Alg.Var 1)):: R_xml_0_rule_as_list_0. Definition R_xml_0_rule_as_list_2 := ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__sharp_ nil)::(algebra.Alg.Var 1)::nil)), (algebra.Alg.Var 1))::R_xml_0_rule_as_list_1. Definition R_xml_0_rule_as_list_3 := ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 2)::nil))::nil)), (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil))::nil))):: R_xml_0_rule_as_list_2. Definition R_xml_0_rule_as_list_4 := ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 2)::nil))::nil)), (algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil))::nil))):: R_xml_0_rule_as_list_3. Definition R_xml_0_rule_as_list_5 := ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 2)::nil))::nil)), (algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil))::nil))):: R_xml_0_rule_as_list_4. Definition R_xml_0_rule_as_list_6 := ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 2)::nil))::nil)), (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::nil))::nil))):: R_xml_0_rule_as_list_5. Definition R_xml_0_rule_as_list_7 := ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Var 2):: (algebra.Alg.Var 3)::nil))::nil)), (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil)):: (algebra.Alg.Var 3)::nil)))::R_xml_0_rule_as_list_6. Definition R_xml_0_rule_as_list_8 := ((algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id__sharp_ nil)::nil)),(algebra.Alg.Var 1)):: R_xml_0_rule_as_list_7. Definition R_xml_0_rule_as_list_9 := ((algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id__sharp_ nil)::(algebra.Alg.Var 1)::nil)), (algebra.Alg.Term algebra.F.id__sharp_ nil))::R_xml_0_rule_as_list_8. Definition R_xml_0_rule_as_list_10 := ((algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 2)::nil))::nil)), (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil))::nil))):: R_xml_0_rule_as_list_9. Definition R_xml_0_rule_as_list_11 := ((algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 2)::nil))::nil)), (algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::nil))::nil))):: R_xml_0_rule_as_list_10. Definition R_xml_0_rule_as_list_12 := ((algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 2)::nil))::nil)), (algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil))::nil))):: R_xml_0_rule_as_list_11. Definition R_xml_0_rule_as_list_13 := ((algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 2)::nil))::nil)), (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil))::nil))):: R_xml_0_rule_as_list_12. Definition R_xml_0_rule_as_list_14 := ((algebra.Alg.Term algebra.F.id_not ((algebra.Alg.Term algebra.F.id_false nil)::nil)),(algebra.Alg.Term algebra.F.id_true nil)):: R_xml_0_rule_as_list_13. Definition R_xml_0_rule_as_list_15 := ((algebra.Alg.Term algebra.F.id_not ((algebra.Alg.Term algebra.F.id_true nil)::nil)),(algebra.Alg.Term algebra.F.id_false nil)):: R_xml_0_rule_as_list_14. Definition R_xml_0_rule_as_list_16 := ((algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id_true nil)::nil)),(algebra.Alg.Var 1)):: R_xml_0_rule_as_list_15. Definition R_xml_0_rule_as_list_17 := ((algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id_false nil)::nil)), (algebra.Alg.Term algebra.F.id_false nil))::R_xml_0_rule_as_list_16. Definition R_xml_0_rule_as_list_18 := ((algebra.Alg.Term algebra.F.id_if ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil)), (algebra.Alg.Var 1))::R_xml_0_rule_as_list_17. Definition R_xml_0_rule_as_list_19 := ((algebra.Alg.Term algebra.F.id_if ((algebra.Alg.Term algebra.F.id_false nil)::(algebra.Alg.Var 1)::(algebra.Alg.Var 2)::nil)), (algebra.Alg.Var 2))::R_xml_0_rule_as_list_18. Definition R_xml_0_rule_as_list_20 := ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 2)::nil))::nil)), (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil)))::R_xml_0_rule_as_list_19. Definition R_xml_0_rule_as_list_21 := ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 2)::nil))::nil)), (algebra.Alg.Term algebra.F.id_not ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Var 2)::(algebra.Alg.Var 1)::nil))::nil))):: R_xml_0_rule_as_list_20. Definition R_xml_0_rule_as_list_22 := ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 2)::nil))::nil)), (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil)))::R_xml_0_rule_as_list_21. Definition R_xml_0_rule_as_list_23 := ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 2)::nil))::nil)), (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Var 1):: (algebra.Alg.Var 2)::nil)))::R_xml_0_rule_as_list_22. Definition R_xml_0_rule_as_list_24 := ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id__sharp_ nil)::nil)), (algebra.Alg.Term algebra.F.id_true nil))::R_xml_0_rule_as_list_23. Definition R_xml_0_rule_as_list_25 := ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id__sharp_ nil)::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_false nil))::R_xml_0_rule_as_list_24. Definition R_xml_0_rule_as_list_26 := ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id__sharp_ nil)::(algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id__sharp_ nil)::(algebra.Alg.Var 1)::nil))):: R_xml_0_rule_as_list_25. Definition R_xml_0_rule_as_list_27 := ((algebra.Alg.Term algebra.F.id_val ((algebra.Alg.Term algebra.F.id_l ((algebra.Alg.Var 1)::nil))::nil)),(algebra.Alg.Var 1)):: R_xml_0_rule_as_list_26. Definition R_xml_0_rule_as_list_28 := ((algebra.Alg.Term algebra.F.id_val ((algebra.Alg.Term algebra.F.id_n ((algebra.Alg.Var 1)::(algebra.Alg.Var 2):: (algebra.Alg.Var 3)::nil))::nil)),(algebra.Alg.Var 1)):: R_xml_0_rule_as_list_27. Definition R_xml_0_rule_as_list_29 := ((algebra.Alg.Term algebra.F.id_min ((algebra.Alg.Term algebra.F.id_l ((algebra.Alg.Var 1)::nil))::nil)),(algebra.Alg.Var 1)):: R_xml_0_rule_as_list_28. Definition R_xml_0_rule_as_list_30 := ((algebra.Alg.Term algebra.F.id_min ((algebra.Alg.Term algebra.F.id_n ((algebra.Alg.Var 1)::(algebra.Alg.Var 2):: (algebra.Alg.Var 3)::nil))::nil)), (algebra.Alg.Term algebra.F.id_min ((algebra.Alg.Var 2)::nil))):: R_xml_0_rule_as_list_29. Definition R_xml_0_rule_as_list_31 := ((algebra.Alg.Term algebra.F.id_max ((algebra.Alg.Term algebra.F.id_l ((algebra.Alg.Var 1)::nil))::nil)),(algebra.Alg.Var 1)):: R_xml_0_rule_as_list_30. Definition R_xml_0_rule_as_list_32 := ((algebra.Alg.Term algebra.F.id_max ((algebra.Alg.Term algebra.F.id_n ((algebra.Alg.Var 1)::(algebra.Alg.Var 2):: (algebra.Alg.Var 3)::nil))::nil)), (algebra.Alg.Term algebra.F.id_max ((algebra.Alg.Var 3)::nil))):: R_xml_0_rule_as_list_31. Definition R_xml_0_rule_as_list_33 := ((algebra.Alg.Term algebra.F.id_bs ((algebra.Alg.Term algebra.F.id_l ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_true nil))::R_xml_0_rule_as_list_32. Definition R_xml_0_rule_as_list_34 := ((algebra.Alg.Term algebra.F.id_bs ((algebra.Alg.Term algebra.F.id_n ((algebra.Alg.Var 1)::(algebra.Alg.Var 2):: (algebra.Alg.Var 3)::nil))::nil)), (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id_max ((algebra.Alg.Var 2)::nil))::nil)):: (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_min ((algebra.Alg.Var 3)::nil))::(algebra.Alg.Var 1)::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_bs ((algebra.Alg.Var 2)::nil))::(algebra.Alg.Term algebra.F.id_bs ((algebra.Alg.Var 3)::nil))::nil))::nil)))::R_xml_0_rule_as_list_33. Definition R_xml_0_rule_as_list_35 := ((algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Term algebra.F.id_l ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil)))::R_xml_0_rule_as_list_34. Definition R_xml_0_rule_as_list_36 := ((algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Term algebra.F.id_n ((algebra.Alg.Var 1)::(algebra.Alg.Var 2):: (algebra.Alg.Var 3)::nil))::nil)), (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Var 1)::nil))::(algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Var 2)::nil))::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::nil))):: R_xml_0_rule_as_list_35. Definition R_xml_0_rule_as_list_37 := ((algebra.Alg.Term algebra.F.id_wb ((algebra.Alg.Term algebra.F.id_l ((algebra.Alg.Var 1)::nil))::nil)), (algebra.Alg.Term algebra.F.id_true nil))::R_xml_0_rule_as_list_36. Definition R_xml_0_rule_as_list_38 := ((algebra.Alg.Term algebra.F.id_wb ((algebra.Alg.Term algebra.F.id_n ((algebra.Alg.Var 1)::(algebra.Alg.Var 2):: (algebra.Alg.Var 3)::nil))::nil)), (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_if ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Var 2)::nil))::(algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Var 3)::nil))::nil))::(algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::(algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Var 2)::nil))::(algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Var 3)::nil))::nil))::nil))::(algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::(algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Var 3)::nil))::(algebra.Alg.Term algebra.F.id_size ((algebra.Alg.Var 2)::nil))::nil))::nil))::nil))::(algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_wb ((algebra.Alg.Var 2)::nil))::(algebra.Alg.Term algebra.F.id_wb ((algebra.Alg.Var 3)::nil))::nil))::nil)))::R_xml_0_rule_as_list_37. Definition R_xml_0_rule_as_list := R_xml_0_rule_as_list_38. Lemma R_xml_0_rules_included : forall l r, R_xml_0_rules r l <-> In (l,r) R_xml_0_rule_as_list. Proof. intros l r. constructor. intros H. case H;clear H; (apply (more_list.mem_impl_in (@eq (algebra.Alg.term*algebra.Alg.term))); [tauto|idtac]); match goal with | |- _ _ _ ?t ?l => let u := fresh "u" in (generalize (more_list.mem_bool_ok _ _ algebra.Alg_ext.eq_term_term_bool_ok t l); set (u:=more_list.mem_bool algebra.Alg_ext.eq_term_term_bool t l) in *; vm_compute in u|-;unfold u in *;clear u;intros H;refine H) end . intros H. vm_compute in H|-. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 39. injection H;intros ;subst;constructor 38. injection H;intros ;subst;constructor 37. injection H;intros ;subst;constructor 36. injection H;intros ;subst;constructor 35. injection H;intros ;subst;constructor 34. injection H;intros ;subst;constructor 33. injection H;intros ;subst;constructor 32. injection H;intros ;subst;constructor 31. injection H;intros ;subst;constructor 30. injection H;intros ;subst;constructor 29. injection H;intros ;subst;constructor 28. injection H;intros ;subst;constructor 27. injection H;intros ;subst;constructor 26. injection H;intros ;subst;constructor 25. injection H;intros ;subst;constructor 24. injection H;intros ;subst;constructor 23. injection H;intros ;subst;constructor 22. injection H;intros ;subst;constructor 21. injection H;intros ;subst;constructor 20. injection H;intros ;subst;constructor 19. injection H;intros ;subst;constructor 18. injection H;intros ;subst;constructor 17. injection H;intros ;subst;constructor 16. 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_6 (x5 x6 x7 x8 x9 x10:Prop) : Prop := | conj_6 : x5->x6->x7->x8->x9->x10->and_6 x5 x6 x7 x8 x9 x10 . 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_6 (forall x6, t = (algebra.Alg.Term algebra.F.id_l (x6::nil)) -> exists x5, t' = (algebra.Alg.Term algebra.F.id_l (x5::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x5 x6)) (t = (algebra.Alg.Term algebra.F.id_false nil) -> t' = (algebra.Alg.Term algebra.F.id_false nil)) (t = (algebra.Alg.Term algebra.F.id__sharp_ nil) -> t' = (algebra.Alg.Term algebra.F.id__sharp_ nil)) (forall x6 x8 x10, t = (algebra.Alg.Term algebra.F.id_n (x6::x8::x10::nil)) -> exists x5, exists x7, exists x9, t' = (algebra.Alg.Term algebra.F.id_n (x5::x7::x9::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x5 x6)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x7 x8)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x9 x10)) (forall x6, t = (algebra.Alg.Term algebra.F.id_1 (x6::nil)) -> exists x5, t' = (algebra.Alg.Term algebra.F.id_1 (x5::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x5 x6)) (t = (algebra.Alg.Term algebra.F.id_true nil) -> t' = (algebra.Alg.Term algebra.F.id_true 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 x6 H;exists x6;intuition;constructor 1. intros H;intuition;constructor 1. intros H;intuition;constructor 1. intros x6 x8 x10 H;exists x6;exists x8;exists x10;intuition;constructor 1. intros x6 H;exists x6;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_l H_id_false H_id__sharp_ H_id_n H_id_1 H_id_true]. constructor. clear H_id_false H_id__sharp_ H_id_n H_id_1 H_id_true;intros x6 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 x6 |- _ => destruct (H_id_l y (refl_equal _)) as [x5];intros ;intuition;exists x5; intuition;eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_l H_id__sharp_ H_id_n H_id_1 H_id_true;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_l H_id_false H_id_n H_id_1 H_id_true;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_l H_id_false H_id__sharp_ H_id_1 H_id_true;intros x6 x8 x10 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 x6 |- _ => destruct (H_id_n y x8 x10 (refl_equal _)) as [x5 [x7 [x9]]];intros ; intuition;exists x5;exists x7;exists x9;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x8 |- _ => destruct (H_id_n x6 y x10 (refl_equal _)) as [x5 [x7 [x9]]];intros ; intuition;exists x5;exists x7;exists x9;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x10 |- _ => destruct (H_id_n x6 x8 y (refl_equal _)) as [x5 [x7 [x9]]];intros ; intuition;exists x5;exists x7;exists x9;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_l H_id_false H_id__sharp_ H_id_n H_id_true;intros x6 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 x6 |- _ => destruct (H_id_1 y (refl_equal _)) as [x5];intros ;intuition;exists x5; intuition;eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_l H_id_false H_id__sharp_ H_id_n H_id_1;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_l_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x6, t = (algebra.Alg.Term algebra.F.id_l (x6::nil)) -> exists x5, t' = (algebra.Alg.Term algebra.F.id_l (x5::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x5 x6). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. 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__sharp__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__sharp_ nil) -> t' = (algebra.Alg.Term algebra.F.id__sharp_ nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_n_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x6 x8 x10, t = (algebra.Alg.Term algebra.F.id_n (x6::x8::x10::nil)) -> exists x5, exists x7, exists x9, t' = (algebra.Alg.Term algebra.F.id_n (x5::x7::x9::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x5 x6)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x7 x8)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x9 x10). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_1_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x6, t = (algebra.Alg.Term algebra.F.id_1 (x6::nil)) -> exists x5, t' = (algebra.Alg.Term algebra.F.id_1 (x5::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x5 x6). 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. 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_l (?x5::nil)) |- _ => let x5 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_l_is_R_xml_0_constructor H (refl_equal _)) as [x5 [Heq Hred1]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_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__sharp_ nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id__sharp__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_n (?x7::?x6::?x5::nil)) |- _ => let x7 := fresh "x" in (let x6 := fresh "x" in (let x5 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (let Hred3 := fresh "Hred" in (destruct (id_n_is_R_xml_0_constructor H (refl_equal _)) as [x7 [x6 [x5 [Heq [Hred3 [Hred2 Hred1]]]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_1 (?x5::nil)) |- _ => let x5 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_1_is_R_xml_0_constructor H (refl_equal _)) as [x5 [Heq Hred1]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_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 )) 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_l (?x5::nil)) |- _ => let x5 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_l_is_R_xml_0_constructor H (refl_equal _)) as [x5 [Heq Hred1]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_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__sharp_ nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id__sharp__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_n (?x7::?x6::?x5::nil)) |- _ => let x7 := fresh "x" in (let x6 := fresh "x" in (let x5 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (let Hred3 := fresh "Hred" in (destruct (id_n_is_R_xml_0_constructor H (refl_equal _)) as [x7 [x6 [x5 [Heq [Hred3 [Hred2 Hred1]]]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_1 (?x5::nil)) |- _ => let x5 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_1_is_R_xml_0_constructor H (refl_equal _)) as [x5 [Heq Hred1]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_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 ))) 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_0 : A ->A. Hypothesis P_id_bs : A ->A. Hypothesis P_id_and : A ->A ->A. Hypothesis P_id__minus_ : A ->A ->A. Hypothesis P_id_l : A ->A. Hypothesis P_id__plus_ : A ->A ->A. Hypothesis P_id_wb : A ->A. Hypothesis P_id_ge : A ->A ->A. Hypothesis P_id_false : A. Hypothesis P_id_min : A ->A. Hypothesis P_id__sharp_ : A. Hypothesis P_id_size : A ->A. Hypothesis P_id_if : A ->A ->A ->A. Hypothesis P_id_not : A ->A. Hypothesis P_id_n : A ->A ->A ->A. Hypothesis P_id_1 : A ->A. Hypothesis P_id_val : A ->A. Hypothesis P_id_true : A. Hypothesis P_id_max : A ->A. Hypothesis P_id_0_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_0 x6 <= P_id_0 x5. Hypothesis P_id_bs_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_bs x6 <= P_id_bs x5. Hypothesis P_id_and_monotonic : forall x8 x6 x5 x7, (A0 <= x8)/\ (x8 <= x7) -> (A0 <= x6)/\ (x6 <= x5) ->P_id_and x6 x8 <= P_id_and x5 x7. Hypothesis P_id__minus__monotonic : forall x8 x6 x5 x7, (A0 <= x8)/\ (x8 <= x7) -> (A0 <= x6)/\ (x6 <= x5) ->P_id__minus_ x6 x8 <= P_id__minus_ x5 x7. Hypothesis P_id_l_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_l x6 <= P_id_l x5. Hypothesis P_id__plus__monotonic : forall x8 x6 x5 x7, (A0 <= x8)/\ (x8 <= x7) -> (A0 <= x6)/\ (x6 <= x5) ->P_id__plus_ x6 x8 <= P_id__plus_ x5 x7. Hypothesis P_id_wb_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_wb x6 <= P_id_wb x5. Hypothesis P_id_ge_monotonic : forall x8 x6 x5 x7, (A0 <= x8)/\ (x8 <= x7) -> (A0 <= x6)/\ (x6 <= x5) ->P_id_ge x6 x8 <= P_id_ge x5 x7. Hypothesis P_id_min_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_min x6 <= P_id_min x5. Hypothesis P_id_size_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_size x6 <= P_id_size x5. Hypothesis P_id_if_monotonic : forall x8 x10 x6 x9 x5 x7, (A0 <= x10)/\ (x10 <= x9) -> (A0 <= x8)/\ (x8 <= x7) -> (A0 <= x6)/\ (x6 <= x5) ->P_id_if x6 x8 x10 <= P_id_if x5 x7 x9. Hypothesis P_id_not_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_not x6 <= P_id_not x5. Hypothesis P_id_n_monotonic : forall x8 x10 x6 x9 x5 x7, (A0 <= x10)/\ (x10 <= x9) -> (A0 <= x8)/\ (x8 <= x7) -> (A0 <= x6)/\ (x6 <= x5) ->P_id_n x6 x8 x10 <= P_id_n x5 x7 x9. Hypothesis P_id_1_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_1 x6 <= P_id_1 x5. Hypothesis P_id_val_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_val x6 <= P_id_val x5. Hypothesis P_id_max_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_max x6 <= P_id_max x5. Hypothesis P_id_0_bounded : forall x5, (A0 <= x5) ->A0 <= P_id_0 x5. Hypothesis P_id_bs_bounded : forall x5, (A0 <= x5) ->A0 <= P_id_bs x5. Hypothesis P_id_and_bounded : forall x6 x5, (A0 <= x5) ->(A0 <= x6) ->A0 <= P_id_and x6 x5. Hypothesis P_id__minus__bounded : forall x6 x5, (A0 <= x5) ->(A0 <= x6) ->A0 <= P_id__minus_ x6 x5. Hypothesis P_id_l_bounded : forall x5, (A0 <= x5) ->A0 <= P_id_l x5. Hypothesis P_id__plus__bounded : forall x6 x5, (A0 <= x5) ->(A0 <= x6) ->A0 <= P_id__plus_ x6 x5. Hypothesis P_id_wb_bounded : forall x5, (A0 <= x5) ->A0 <= P_id_wb x5. Hypothesis P_id_ge_bounded : forall x6 x5, (A0 <= x5) ->(A0 <= x6) ->A0 <= P_id_ge x6 x5. Hypothesis P_id_false_bounded : A0 <= P_id_false . Hypothesis P_id_min_bounded : forall x5, (A0 <= x5) ->A0 <= P_id_min x5. Hypothesis P_id__sharp__bounded : A0 <= P_id__sharp_ . Hypothesis P_id_size_bounded : forall x5, (A0 <= x5) ->A0 <= P_id_size x5. Hypothesis P_id_if_bounded : forall x6 x5 x7, (A0 <= x5) ->(A0 <= x6) ->(A0 <= x7) ->A0 <= P_id_if x7 x6 x5. Hypothesis P_id_not_bounded : forall x5, (A0 <= x5) ->A0 <= P_id_not x5. Hypothesis P_id_n_bounded : forall x6 x5 x7, (A0 <= x5) ->(A0 <= x6) ->(A0 <= x7) ->A0 <= P_id_n x7 x6 x5. Hypothesis P_id_1_bounded : forall x5, (A0 <= x5) ->A0 <= P_id_1 x5. Hypothesis P_id_val_bounded : forall x5, (A0 <= x5) ->A0 <= P_id_val x5. Hypothesis P_id_true_bounded : A0 <= P_id_true . Hypothesis P_id_max_bounded : forall x5, (A0 <= x5) ->A0 <= P_id_max x5. Fixpoint measure t { struct t } := match t with | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0 (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_bs (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_and (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) => P_id__minus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_l (x5::nil)) => P_id_l (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) => P_id__plus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_wb (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_ge (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_min (measure x5) | (algebra.Alg.Term algebra.F.id__sharp_ nil) => P_id__sharp_ | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_size (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6::x5::nil)) => P_id_if (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_not (measure x5) | (algebra.Alg.Term algebra.F.id_n (x7::x6::x5::nil)) => P_id_n (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_1 (x5::nil)) => P_id_1 (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_val (measure x5) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_max (measure x5) | _ => A0 end. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0 (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_bs (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_and (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) => P_id__minus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_l (x5::nil)) => P_id_l (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) => P_id__plus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_wb (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_ge (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_min (measure x5) | (algebra.Alg.Term algebra.F.id__sharp_ nil) => P_id__sharp_ | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_size (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6::x5::nil)) => P_id_if (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_not (measure x5) | (algebra.Alg.Term algebra.F.id_n (x7::x6::x5::nil)) => P_id_n (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_1 (x5::nil)) => P_id_1 (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_val (measure x5) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_max (measure x5) | _ => 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_0 => P_id_0 | algebra.F.id_bs => P_id_bs | algebra.F.id_and => P_id_and | algebra.F.id__minus_ => P_id__minus_ | algebra.F.id_l => P_id_l | algebra.F.id__plus_ => P_id__plus_ | algebra.F.id_wb => P_id_wb | algebra.F.id_ge => P_id_ge | algebra.F.id_false => P_id_false | algebra.F.id_min => P_id_min | algebra.F.id__sharp_ => P_id__sharp_ | algebra.F.id_size => P_id_size | algebra.F.id_if => P_id_if | algebra.F.id_not => P_id_not | algebra.F.id_n => P_id_n | algebra.F.id_1 => P_id_1 | algebra.F.id_val => P_id_val | algebra.F.id_true => P_id_true | algebra.F.id_max => P_id_max 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_0 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_bs => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_and => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id__minus_ => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_l => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id__plus_ => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_wb => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_ge => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_false => match l with | nil => _ | _::_ => _ end | algebra.F.id_min => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id__sharp_ => match l with | nil => _ | _::_ => _ end | algebra.F.id_size => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_if => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_not => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_n => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_1 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_val => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_true => match l with | nil => _ | _::_ => _ end | algebra.F.id_max => match l with | nil => _ | _::nil => _ | _::_::_ => _ end end;simpl in |-*;unfold eq_rect_r, eq_rect, sym_eq in |-*; try (reflexivity );f_equal ;auto. Qed. Lemma measure_bounded : forall t, A0 <= measure t. Proof. intros t. rewrite same_measure in |-*. apply (InterpGen.measure_bounded Aop). intros f. case f. vm_compute in |-*;intros ;apply P_id_0_bounded;assumption. vm_compute in |-*;intros ;apply P_id_bs_bounded;assumption. vm_compute in |-*;intros ;apply P_id_and_bounded;assumption. vm_compute in |-*;intros ;apply P_id__minus__bounded;assumption. vm_compute in |-*;intros ;apply P_id_l_bounded;assumption. vm_compute in |-*;intros ;apply P_id__plus__bounded;assumption. vm_compute in |-*;intros ;apply P_id_wb_bounded;assumption. vm_compute in |-*;intros ;apply P_id_ge_bounded;assumption. vm_compute in |-*;intros ;apply P_id_false_bounded;assumption. vm_compute in |-*;intros ;apply P_id_min_bounded;assumption. vm_compute in |-*;intros ;apply P_id__sharp__bounded;assumption. vm_compute in |-*;intros ;apply P_id_size_bounded;assumption. vm_compute in |-*;intros ;apply P_id_if_bounded;assumption. vm_compute in |-*;intros ;apply P_id_not_bounded;assumption. vm_compute in |-*;intros ;apply P_id_n_bounded;assumption. vm_compute in |-*;intros ;apply P_id_1_bounded;assumption. vm_compute in |-*;intros ;apply P_id_val_bounded;assumption. vm_compute in |-*;intros ;apply P_id_true_bounded;assumption. vm_compute in |-*;intros ;apply P_id_max_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_0_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_bs_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_and_monotonic;assumption. vm_compute in |-*;intros ;apply P_id__minus__monotonic;assumption. vm_compute in |-*;intros ;apply P_id_l_monotonic;assumption. vm_compute in |-*;intros ;apply P_id__plus__monotonic;assumption. vm_compute in |-*;intros ;apply P_id_wb_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_ge_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_min_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_size_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_if_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_not_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_n_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_1_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_val_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_max_monotonic;assumption. intros f. case f. vm_compute in |-*;intros ;apply P_id_0_bounded;assumption. vm_compute in |-*;intros ;apply P_id_bs_bounded;assumption. vm_compute in |-*;intros ;apply P_id_and_bounded;assumption. vm_compute in |-*;intros ;apply P_id__minus__bounded;assumption. vm_compute in |-*;intros ;apply P_id_l_bounded;assumption. vm_compute in |-*;intros ;apply P_id__plus__bounded;assumption. vm_compute in |-*;intros ;apply P_id_wb_bounded;assumption. vm_compute in |-*;intros ;apply P_id_ge_bounded;assumption. vm_compute in |-*;intros ;apply P_id_false_bounded;assumption. vm_compute in |-*;intros ;apply P_id_min_bounded;assumption. vm_compute in |-*;intros ;apply P_id__sharp__bounded;assumption. vm_compute in |-*;intros ;apply P_id_size_bounded;assumption. vm_compute in |-*;intros ;apply P_id_if_bounded;assumption. vm_compute in |-*;intros ;apply P_id_not_bounded;assumption. vm_compute in |-*;intros ;apply P_id_n_bounded;assumption. vm_compute in |-*;intros ;apply P_id_1_bounded;assumption. vm_compute in |-*;intros ;apply P_id_val_bounded;assumption. vm_compute in |-*;intros ;apply P_id_true_bounded;assumption. vm_compute in |-*;intros ;apply P_id_max_bounded;assumption. intros . do 2 (rewrite <- same_measure in |-*). apply rules_monotonic;assumption. assumption. Qed. Hypothesis P_id_MIN : A ->A. Hypothesis P_id_0_hat_1 : A ->A. Hypothesis P_id_SIZE : A ->A. Hypothesis P_id_BS : A ->A. Hypothesis P_id_GE : A ->A ->A. Hypothesis P_id_IF : A ->A ->A ->A. Hypothesis P_id_MAX : A ->A. Hypothesis P_id__minus__hat_1 : A ->A ->A. Hypothesis P_id_WB : A ->A. Hypothesis P_id__plus__hat_1 : A ->A ->A. Hypothesis P_id_AND : A ->A ->A. Hypothesis P_id_NOT : A ->A. Hypothesis P_id_Marked_val : A ->A. Hypothesis P_id_MIN_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_MIN x6 <= P_id_MIN x5. Hypothesis P_id_0_hat_1_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_0_hat_1 x6 <= P_id_0_hat_1 x5. Hypothesis P_id_SIZE_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_SIZE x6 <= P_id_SIZE x5. Hypothesis P_id_BS_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_BS x6 <= P_id_BS x5. Hypothesis P_id_GE_monotonic : forall x8 x6 x5 x7, (A0 <= x8)/\ (x8 <= x7) -> (A0 <= x6)/\ (x6 <= x5) ->P_id_GE x6 x8 <= P_id_GE x5 x7. Hypothesis P_id_IF_monotonic : forall x8 x10 x6 x9 x5 x7, (A0 <= x10)/\ (x10 <= x9) -> (A0 <= x8)/\ (x8 <= x7) -> (A0 <= x6)/\ (x6 <= x5) ->P_id_IF x6 x8 x10 <= P_id_IF x5 x7 x9. Hypothesis P_id_MAX_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_MAX x6 <= P_id_MAX x5. Hypothesis P_id__minus__hat_1_monotonic : forall x8 x6 x5 x7, (A0 <= x8)/\ (x8 <= x7) -> (A0 <= x6)/\ (x6 <= x5) -> P_id__minus__hat_1 x6 x8 <= P_id__minus__hat_1 x5 x7. Hypothesis P_id_WB_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_WB x6 <= P_id_WB x5. Hypothesis P_id__plus__hat_1_monotonic : forall x8 x6 x5 x7, (A0 <= x8)/\ (x8 <= x7) -> (A0 <= x6)/\ (x6 <= x5) -> P_id__plus__hat_1 x6 x8 <= P_id__plus__hat_1 x5 x7. Hypothesis P_id_AND_monotonic : forall x8 x6 x5 x7, (A0 <= x8)/\ (x8 <= x7) -> (A0 <= x6)/\ (x6 <= x5) ->P_id_AND x6 x8 <= P_id_AND x5 x7. Hypothesis P_id_NOT_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_NOT x6 <= P_id_NOT x5. Hypothesis P_id_Marked_val_monotonic : forall x6 x5, (A0 <= x6)/\ (x6 <= x5) ->P_id_Marked_val x6 <= P_id_Marked_val x5. Definition marked_measure t := match t with | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_MIN (measure x5) | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0_hat_1 (measure x5) | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_SIZE (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_BS (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_GE (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6::x5::nil)) => P_id_IF (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_MAX (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) => P_id__minus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_WB (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) => P_id__plus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_AND (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_NOT (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_Marked_val (measure x5) | _ => 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 x5;apply (P_id_MAX x5). 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 x5;apply (P_id_Marked_val x5). 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 x5;apply (P_id_NOT x5). 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 x7 x6 x5;apply (P_id_IF x7 x6 x5). 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 x5;apply (P_id_SIZE x5). 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 x5;apply (P_id_MIN x5). 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 x6 x5;apply (P_id_GE x6 x5). 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 x5;apply (P_id_WB x5). 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 x6 x5;apply (P_id__plus__hat_1 x6 x5). 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 x6 x5;apply (P_id__minus__hat_1 x6 x5). 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 x6 x5;apply (P_id_AND x6 x5). 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 x5;apply (P_id_BS x5). 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 x5;apply (P_id_0_hat_1 x5). 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_0 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_bs => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_and => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id__minus_ => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_l => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id__plus_ => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_wb => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_ge => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_false => match l with | nil => _ | _::_ => _ end | algebra.F.id_min => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id__sharp_ => match l with | nil => _ | _::_ => _ end | algebra.F.id_size => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_if => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_not => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_n => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_1 => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_val => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_true => match l with | nil => _ | _::_ => _ end | algebra.F.id_max => match l with | nil => _ | _::nil => _ | _::_::_ => _ end end. vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . Qed. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_MIN (measure x5) | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0_hat_1 (measure x5) | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_SIZE (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_BS (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_GE (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6:: x5::nil)) => P_id_IF (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_MAX (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6:: x5::nil)) => P_id__minus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_WB (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6:: x5::nil)) => P_id__plus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_AND (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_NOT (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_Marked_val (measure x5) | _ => 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_0_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_bs_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_and_monotonic;assumption. vm_compute in |-*;intros ;apply P_id__minus__monotonic;assumption. vm_compute in |-*;intros ;apply P_id_l_monotonic;assumption. vm_compute in |-*;intros ;apply P_id__plus__monotonic;assumption. vm_compute in |-*;intros ;apply P_id_wb_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_ge_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_min_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_size_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_if_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_not_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_n_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_1_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_val_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_max_monotonic;assumption. clear f. intros f. case f. vm_compute in |-*;intros ;apply P_id_0_bounded;assumption. vm_compute in |-*;intros ;apply P_id_bs_bounded;assumption. vm_compute in |-*;intros ;apply P_id_and_bounded;assumption. vm_compute in |-*;intros ;apply P_id__minus__bounded;assumption. vm_compute in |-*;intros ;apply P_id_l_bounded;assumption. vm_compute in |-*;intros ;apply P_id__plus__bounded;assumption. vm_compute in |-*;intros ;apply P_id_wb_bounded;assumption. vm_compute in |-*;intros ;apply P_id_ge_bounded;assumption. vm_compute in |-*;intros ;apply P_id_false_bounded;assumption. vm_compute in |-*;intros ;apply P_id_min_bounded;assumption. vm_compute in |-*;intros ;apply P_id__sharp__bounded;assumption. vm_compute in |-*;intros ;apply P_id_size_bounded;assumption. vm_compute in |-*;intros ;apply P_id_if_bounded;assumption. vm_compute in |-*;intros ;apply P_id_not_bounded;assumption. vm_compute in |-*;intros ;apply P_id_n_bounded;assumption. vm_compute in |-*;intros ;apply P_id_1_bounded;assumption. vm_compute in |-*;intros ;apply P_id_val_bounded;assumption. vm_compute in |-*;intros ;apply P_id_true_bounded;assumption. vm_compute in |-*;intros ;apply P_id_max_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_MAX_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_Marked_val_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_NOT_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_IF_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_SIZE_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_MIN_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_GE_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_WB_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__plus__hat_1_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id__minus__hat_1_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_AND_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_BS_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_0_hat_1_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_0 : Z ->Z. Hypothesis P_id_bs : Z ->Z. Hypothesis P_id_and : Z ->Z ->Z. Hypothesis P_id__minus_ : Z ->Z ->Z. Hypothesis P_id_l : Z ->Z. Hypothesis P_id__plus_ : Z ->Z ->Z. Hypothesis P_id_wb : Z ->Z. Hypothesis P_id_ge : Z ->Z ->Z. Hypothesis P_id_false : Z. Hypothesis P_id_min : Z ->Z. Hypothesis P_id__sharp_ : Z. Hypothesis P_id_size : Z ->Z. Hypothesis P_id_if : Z ->Z ->Z ->Z. Hypothesis P_id_not : Z ->Z. Hypothesis P_id_n : Z ->Z ->Z ->Z. Hypothesis P_id_1 : Z ->Z. Hypothesis P_id_val : Z ->Z. Hypothesis P_id_true : Z. Hypothesis P_id_max : Z ->Z. Hypothesis P_id_0_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_0 x6 <= P_id_0 x5. Hypothesis P_id_bs_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_bs x6 <= P_id_bs x5. Hypothesis P_id_and_monotonic : forall x8 x6 x5 x7, (min_value <= x8)/\ (x8 <= x7) -> (min_value <= x6)/\ (x6 <= x5) ->P_id_and x6 x8 <= P_id_and x5 x7. Hypothesis P_id__minus__monotonic : forall x8 x6 x5 x7, (min_value <= x8)/\ (x8 <= x7) -> (min_value <= x6)/\ (x6 <= x5) -> P_id__minus_ x6 x8 <= P_id__minus_ x5 x7. Hypothesis P_id_l_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_l x6 <= P_id_l x5. Hypothesis P_id__plus__monotonic : forall x8 x6 x5 x7, (min_value <= x8)/\ (x8 <= x7) -> (min_value <= x6)/\ (x6 <= x5) ->P_id__plus_ x6 x8 <= P_id__plus_ x5 x7. Hypothesis P_id_wb_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_wb x6 <= P_id_wb x5. Hypothesis P_id_ge_monotonic : forall x8 x6 x5 x7, (min_value <= x8)/\ (x8 <= x7) -> (min_value <= x6)/\ (x6 <= x5) ->P_id_ge x6 x8 <= P_id_ge x5 x7. Hypothesis P_id_min_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_min x6 <= P_id_min x5. Hypothesis P_id_size_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_size x6 <= P_id_size x5. Hypothesis P_id_if_monotonic : forall x8 x10 x6 x9 x5 x7, (min_value <= x10)/\ (x10 <= x9) -> (min_value <= x8)/\ (x8 <= x7) -> (min_value <= x6)/\ (x6 <= x5) ->P_id_if x6 x8 x10 <= P_id_if x5 x7 x9. Hypothesis P_id_not_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_not x6 <= P_id_not x5. Hypothesis P_id_n_monotonic : forall x8 x10 x6 x9 x5 x7, (min_value <= x10)/\ (x10 <= x9) -> (min_value <= x8)/\ (x8 <= x7) -> (min_value <= x6)/\ (x6 <= x5) ->P_id_n x6 x8 x10 <= P_id_n x5 x7 x9. Hypothesis P_id_1_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_1 x6 <= P_id_1 x5. Hypothesis P_id_val_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_val x6 <= P_id_val x5. Hypothesis P_id_max_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_max x6 <= P_id_max x5. Hypothesis P_id_0_bounded : forall x5, (min_value <= x5) ->min_value <= P_id_0 x5. Hypothesis P_id_bs_bounded : forall x5, (min_value <= x5) ->min_value <= P_id_bs x5. Hypothesis P_id_and_bounded : forall x6 x5, (min_value <= x5) ->(min_value <= x6) ->min_value <= P_id_and x6 x5. Hypothesis P_id__minus__bounded : forall x6 x5, (min_value <= x5) ->(min_value <= x6) ->min_value <= P_id__minus_ x6 x5. Hypothesis P_id_l_bounded : forall x5, (min_value <= x5) ->min_value <= P_id_l x5. Hypothesis P_id__plus__bounded : forall x6 x5, (min_value <= x5) ->(min_value <= x6) ->min_value <= P_id__plus_ x6 x5. Hypothesis P_id_wb_bounded : forall x5, (min_value <= x5) ->min_value <= P_id_wb x5. Hypothesis P_id_ge_bounded : forall x6 x5, (min_value <= x5) ->(min_value <= x6) ->min_value <= P_id_ge x6 x5. Hypothesis P_id_false_bounded : min_value <= P_id_false . Hypothesis P_id_min_bounded : forall x5, (min_value <= x5) ->min_value <= P_id_min x5. Hypothesis P_id__sharp__bounded : min_value <= P_id__sharp_ . Hypothesis P_id_size_bounded : forall x5, (min_value <= x5) ->min_value <= P_id_size x5. Hypothesis P_id_if_bounded : forall x6 x5 x7, (min_value <= x5) -> (min_value <= x6) ->(min_value <= x7) ->min_value <= P_id_if x7 x6 x5. Hypothesis P_id_not_bounded : forall x5, (min_value <= x5) ->min_value <= P_id_not x5. Hypothesis P_id_n_bounded : forall x6 x5 x7, (min_value <= x5) -> (min_value <= x6) ->(min_value <= x7) ->min_value <= P_id_n x7 x6 x5. Hypothesis P_id_1_bounded : forall x5, (min_value <= x5) ->min_value <= P_id_1 x5. Hypothesis P_id_val_bounded : forall x5, (min_value <= x5) ->min_value <= P_id_val x5. Hypothesis P_id_true_bounded : min_value <= P_id_true . Hypothesis P_id_max_bounded : forall x5, (min_value <= x5) ->min_value <= P_id_max x5. Definition measure := Interp.measure min_value P_id_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0 (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_bs (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_and (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) => P_id__minus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_l (x5::nil)) => P_id_l (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) => P_id__plus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_wb (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_ge (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_min (measure x5) | (algebra.Alg.Term algebra.F.id__sharp_ nil) => P_id__sharp_ | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_size (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6::x5::nil)) => P_id_if (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_not (measure x5) | (algebra.Alg.Term algebra.F.id_n (x7::x6::x5::nil)) => P_id_n (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_1 (x5::nil)) => P_id_1 (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_val (measure x5) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_max (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. Qed. Hypothesis P_id_MIN : Z ->Z. Hypothesis P_id_0_hat_1 : Z ->Z. Hypothesis P_id_SIZE : Z ->Z. Hypothesis P_id_BS : Z ->Z. Hypothesis P_id_GE : Z ->Z ->Z. Hypothesis P_id_IF : Z ->Z ->Z ->Z. Hypothesis P_id_MAX : Z ->Z. Hypothesis P_id__minus__hat_1 : Z ->Z ->Z. Hypothesis P_id_WB : Z ->Z. Hypothesis P_id__plus__hat_1 : Z ->Z ->Z. Hypothesis P_id_AND : Z ->Z ->Z. Hypothesis P_id_NOT : Z ->Z. Hypothesis P_id_Marked_val : Z ->Z. Hypothesis P_id_MIN_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_MIN x6 <= P_id_MIN x5. Hypothesis P_id_0_hat_1_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_0_hat_1 x6 <= P_id_0_hat_1 x5. Hypothesis P_id_SIZE_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_SIZE x6 <= P_id_SIZE x5. Hypothesis P_id_BS_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_BS x6 <= P_id_BS x5. Hypothesis P_id_GE_monotonic : forall x8 x6 x5 x7, (min_value <= x8)/\ (x8 <= x7) -> (min_value <= x6)/\ (x6 <= x5) ->P_id_GE x6 x8 <= P_id_GE x5 x7. Hypothesis P_id_IF_monotonic : forall x8 x10 x6 x9 x5 x7, (min_value <= x10)/\ (x10 <= x9) -> (min_value <= x8)/\ (x8 <= x7) -> (min_value <= x6)/\ (x6 <= x5) ->P_id_IF x6 x8 x10 <= P_id_IF x5 x7 x9. Hypothesis P_id_MAX_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_MAX x6 <= P_id_MAX x5. Hypothesis P_id__minus__hat_1_monotonic : forall x8 x6 x5 x7, (min_value <= x8)/\ (x8 <= x7) -> (min_value <= x6)/\ (x6 <= x5) -> P_id__minus__hat_1 x6 x8 <= P_id__minus__hat_1 x5 x7. Hypothesis P_id_WB_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_WB x6 <= P_id_WB x5. Hypothesis P_id__plus__hat_1_monotonic : forall x8 x6 x5 x7, (min_value <= x8)/\ (x8 <= x7) -> (min_value <= x6)/\ (x6 <= x5) -> P_id__plus__hat_1 x6 x8 <= P_id__plus__hat_1 x5 x7. Hypothesis P_id_AND_monotonic : forall x8 x6 x5 x7, (min_value <= x8)/\ (x8 <= x7) -> (min_value <= x6)/\ (x6 <= x5) ->P_id_AND x6 x8 <= P_id_AND x5 x7. Hypothesis P_id_NOT_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) ->P_id_NOT x6 <= P_id_NOT x5. Hypothesis P_id_Marked_val_monotonic : forall x6 x5, (min_value <= x6)/\ (x6 <= x5) -> P_id_Marked_val x6 <= P_id_Marked_val x5. Definition marked_measure := Interp.marked_measure min_value P_id_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max P_id_MIN P_id_0_hat_1 P_id_SIZE P_id_BS P_id_GE P_id_IF P_id_MAX P_id__minus__hat_1 P_id_WB P_id__plus__hat_1 P_id_AND P_id_NOT P_id_Marked_val. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_MIN (measure x5) | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0_hat_1 (measure x5) | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_SIZE (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_BS (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_GE (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6:: x5::nil)) => P_id_IF (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_MAX (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6:: x5::nil)) => P_id__minus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_WB (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6:: x5::nil)) => P_id__plus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_AND (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_NOT (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_Marked_val (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. intros ;apply P_id_MIN_monotonic;assumption. intros ;apply P_id_0_hat_1_monotonic;assumption. intros ;apply P_id_SIZE_monotonic;assumption. intros ;apply P_id_BS_monotonic;assumption. intros ;apply P_id_GE_monotonic;assumption. intros ;apply P_id_IF_monotonic;assumption. intros ;apply P_id_MAX_monotonic;assumption. intros ;apply P_id__minus__hat_1_monotonic;assumption. intros ;apply P_id_WB_monotonic;assumption. intros ;apply P_id__plus__hat_1_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_NOT_monotonic;assumption. intros ;apply P_id_Marked_val_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 := (* <+(0(x_),0(y_)),0(+(x_,y_))> *) | DP_R_xml_0_0 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__plus_ (x1::x2::nil))::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(0(x_),0(y_)),+(x_,y_)> *) | DP_R_xml_0_1 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id__plus_ (x1::x2::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(0(x_),1(y_)),+(x_,y_)> *) | DP_R_xml_0_2 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id__plus_ (x1::x2::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(1(x_),0(y_)),+(x_,y_)> *) | DP_R_xml_0_3 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id__plus_ (x1::x2::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(1(x_),1(y_)),0(+(+(x_,y_),1(#)))> *) | DP_R_xml_0_4 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ (x1::x2::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(1(x_),1(y_)),+(+(x_,y_),1(#))> *) | DP_R_xml_0_5 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ (x1::x2::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(1(x_),1(y_)),+(x_,y_)> *) | DP_R_xml_0_6 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id__plus_ (x1::x2::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(x_,+(y_,z_)),+(+(x_,y_),z_)> *) | DP_R_xml_0_7 : forall x2 x6 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id__plus_ (x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ (x1::x2::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(x_,+(y_,z_)),+(x_,y_)> *) | DP_R_xml_0_8 : forall x2 x6 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id__plus_ (x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id__plus_ (x1::x2::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <-(0(x_),0(y_)),0(-(x_,y_))> *) | DP_R_xml_0_9 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__minus_ (x1::x2::nil))::nil)) (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) (* <-(0(x_),0(y_)),-(x_,y_)> *) | DP_R_xml_0_10 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id__minus_ (x1::x2::nil)) (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) (* <-(0(x_),1(y_)),-(-(x_,y_),1(#))> *) | DP_R_xml_0_11 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id__minus_ (x1::x2::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) (* <-(0(x_),1(y_)),-(x_,y_)> *) | DP_R_xml_0_12 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id__minus_ (x1::x2::nil)) (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) (* <-(1(x_),0(y_)),-(x_,y_)> *) | DP_R_xml_0_13 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id__minus_ (x1::x2::nil)) (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) (* <-(1(x_),1(y_)),0(-(x_,y_))> *) | DP_R_xml_0_14 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__minus_ (x1::x2::nil))::nil)) (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) (* <-(1(x_),1(y_)),-(x_,y_)> *) | DP_R_xml_0_15 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id__minus_ (x1::x2::nil)) (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) (* *) | DP_R_xml_0_16 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_ge (x1::x2::nil)) (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) (* *) | DP_R_xml_0_17 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_not ((algebra.Alg.Term algebra.F.id_ge (x2::x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) (* *) | DP_R_xml_0_18 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_ge (x2::x1::nil)) (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) (* *) | DP_R_xml_0_19 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_ge (x1::x2::nil)) (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) (* *) | DP_R_xml_0_20 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_ge (x1::x2::nil)) (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) (* *) | DP_R_xml_0_21 : forall x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id__sharp_ nil) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id__sharp_ nil)::x1::nil)) (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) (* *) | DP_R_xml_0_22 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_min (x2::nil)) (algebra.Alg.Term algebra.F.id_min (x5::nil)) (* *) | DP_R_xml_0_23 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_max (x3::nil)) (algebra.Alg.Term algebra.F.id_max (x5::nil)) (* *) | DP_R_xml_0_24 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_ge (x1:: (algebra.Alg.Term algebra.F.id_max (x2::nil))::nil)):: (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_min (x3::nil))::x1::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_bs (x2::nil))::(algebra.Alg.Term algebra.F.id_bs (x3::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) (* *) | DP_R_xml_0_25 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_ge (x1::(algebra.Alg.Term algebra.F.id_max (x2::nil))::nil))::(algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_min (x3::nil)):: x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) (* *) | DP_R_xml_0_26 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_ge (x1::(algebra.Alg.Term algebra.F.id_max (x2::nil))::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) (* *) | DP_R_xml_0_27 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_max (x2::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) (* *) | DP_R_xml_0_28 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_min (x3::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) (* *) | DP_R_xml_0_29 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_min (x3::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) (* *) | DP_R_xml_0_30 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_bs (x2::nil))::(algebra.Alg.Term algebra.F.id_bs (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) (* *) | DP_R_xml_0_31 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_bs (x2::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) (* *) | DP_R_xml_0_32 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_bs (x3::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) (* *) | DP_R_xml_0_33 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id_size (x1::nil))::(algebra.Alg.Term algebra.F.id_size (x2::nil))::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_size (x5::nil)) (* *) | DP_R_xml_0_34 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id_size (x1::nil))::(algebra.Alg.Term algebra.F.id_size (x2::nil))::nil)) (algebra.Alg.Term algebra.F.id_size (x5::nil)) (* *) | DP_R_xml_0_35 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_size (x1::nil)) (algebra.Alg.Term algebra.F.id_size (x5::nil)) (* *) | DP_R_xml_0_36 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_size (x2::nil)) (algebra.Alg.Term algebra.F.id_size (x5::nil)) (* *) | DP_R_xml_0_37 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_if ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_size (x2::nil)):: (algebra.Alg.Term algebra.F.id_size (x3::nil))::nil)):: (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::(algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size (x2::nil)):: (algebra.Alg.Term algebra.F.id_size (x3::nil))::nil))::nil))::(algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil)):: (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size (x3::nil))::(algebra.Alg.Term algebra.F.id_size (x2::nil))::nil))::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_wb (x2::nil))::(algebra.Alg.Term algebra.F.id_wb (x3::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) (* *) | DP_R_xml_0_38 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_if ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_size (x2::nil))::(algebra.Alg.Term algebra.F.id_size (x3::nil))::nil))::(algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::(algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size (x2::nil))::(algebra.Alg.Term algebra.F.id_size (x3::nil))::nil))::nil))::(algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil)):: (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size (x3::nil))::(algebra.Alg.Term algebra.F.id_size (x2::nil))::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) (* *) | DP_R_xml_0_39 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_size (x2::nil))::(algebra.Alg.Term algebra.F.id_size (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) (* *) | DP_R_xml_0_40 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_size (x2::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) (* *) | DP_R_xml_0_41 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_size (x3::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) (* *) | DP_R_xml_0_42 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::(algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size (x2::nil)):: (algebra.Alg.Term algebra.F.id_size (x3::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) (* *) | DP_R_xml_0_43 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size (x2::nil))::(algebra.Alg.Term algebra.F.id_size (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) (* *) | DP_R_xml_0_44 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::(algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size (x3::nil)):: (algebra.Alg.Term algebra.F.id_size (x2::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) (* *) | DP_R_xml_0_45 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size (x3::nil))::(algebra.Alg.Term algebra.F.id_size (x2::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) (* *) | DP_R_xml_0_46 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_wb (x2::nil))::(algebra.Alg.Term algebra.F.id_wb (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) (* *) | DP_R_xml_0_47 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_wb (x2::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) (* *) | DP_R_xml_0_48 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_wb (x3::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::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 x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_1 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_wb (x2::nil)):: (algebra.Alg.Term algebra.F.id_wb (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::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 x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_2 (algebra.Alg.Term algebra.F.id_if ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_size (x2::nil))::(algebra.Alg.Term algebra.F.id_size (x3::nil))::nil))::(algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil)):: (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size (x2::nil))::(algebra.Alg.Term algebra.F.id_size (x3::nil))::nil))::nil))::(algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil)):: (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size (x3::nil))::(algebra.Alg.Term algebra.F.id_size (x2::nil))::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::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 x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_3 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_if ((algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_size (x2::nil))::(algebra.Alg.Term algebra.F.id_size (x3::nil))::nil))::(algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil)):: (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size (x2::nil))::(algebra.Alg.Term algebra.F.id_size (x3::nil))::nil))::nil))::(algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil)):: (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size (x3::nil))::(algebra.Alg.Term algebra.F.id_size (x2::nil))::nil))::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_wb (x2::nil)):: (algebra.Alg.Term algebra.F.id_wb (x3::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_3 : forall x y, (DP_R_xml_0_non_scc_3 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_4 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_4_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_4 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_bs (x2::nil)):: (algebra.Alg.Term algebra.F.id_bs (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_4 : forall x y, (DP_R_xml_0_non_scc_4 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_5 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_5_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_5 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_ge (x1:: (algebra.Alg.Term algebra.F.id_max (x2::nil))::nil))::(algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_min (x3::nil))::x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_5 : forall x y, (DP_R_xml_0_non_scc_5 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_6 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_6_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_6 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_ge (x1:: (algebra.Alg.Term algebra.F.id_max (x2::nil))::nil))::(algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_min (x3::nil))::x1::nil))::nil)):: (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_bs (x2::nil)):: (algebra.Alg.Term algebra.F.id_bs (x3::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_6 : forall x y, (DP_R_xml_0_non_scc_6 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_scc_7 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_7_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_scc_7 (algebra.Alg.Term algebra.F.id_max (x3::nil)) (algebra.Alg.Term algebra.F.id_max (x5::nil)) . Module WF_DP_R_xml_0_scc_7. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_0 (x5:Z) := 0. Definition P_id_bs (x5:Z) := 0. Definition P_id_and (x5:Z) (x6:Z) := 2* x5. Definition P_id__minus_ (x5:Z) (x6:Z) := 2* x5. Definition P_id_l (x5:Z) := 1* x5. Definition P_id__plus_ (x5:Z) (x6:Z) := 1* x5 + 2* x6. Definition P_id_wb (x5:Z) := 0. Definition P_id_ge (x5:Z) (x6:Z) := 0. Definition P_id_false := 0. Definition P_id_min (x5:Z) := 1* x5. Definition P_id__sharp_ := 0. Definition P_id_size (x5:Z) := 0. Definition P_id_if (x5:Z) (x6:Z) (x7:Z) := 2* x6 + 2* x7. Definition P_id_not (x5:Z) := 0. Definition P_id_n (x5:Z) (x6:Z) (x7:Z) := 1 + 1* x5 + 1* x6 + 3* x7. Definition P_id_1 (x5:Z) := 0. Definition P_id_val (x5:Z) := 1* x5. Definition P_id_true := 0. Definition P_id_max (x5:Z) := 3* x5. Lemma P_id_0_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0 x6 <= P_id_0 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_bs x6 <= P_id_bs x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_and x6 x8 <= P_id_and x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__minus_ x6 x8 <= P_id__minus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_l x6 <= P_id_l x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__plus_ x6 x8 <= P_id__plus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_wb x6 <= P_id_wb x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_ge x6 x8 <= P_id_ge x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_min_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_min x6 <= P_id_min x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_size x6 <= P_id_size x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_if x6 x8 x10 <= P_id_if x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_not_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_not x6 <= P_id_not x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_n x6 x8 x10 <= P_id_n x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_1 x6 <= P_id_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_val x6 <= P_id_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_max_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_max x6 <= P_id_max x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : forall x5, (0 <= x5) ->0 <= P_id_0 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_bounded : forall x5, (0 <= x5) ->0 <= P_id_bs x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_and x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__minus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_bounded : forall x5, (0 <= x5) ->0 <= P_id_l x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__plus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_bounded : forall x5, (0 <= x5) ->0 <= P_id_wb x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_ge x6 x5. 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_min_bounded : forall x5, (0 <= x5) ->0 <= P_id_min x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__sharp__bounded : 0 <= P_id__sharp_ . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_bounded : forall x5, (0 <= x5) ->0 <= P_id_size x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_if x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_not_bounded : forall x5, (0 <= x5) ->0 <= P_id_not x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_n x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_1_bounded : forall x5, (0 <= x5) ->0 <= P_id_1 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_bounded : forall x5, (0 <= x5) ->0 <= P_id_val x5. 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_max_bounded : forall x5, (0 <= x5) ->0 <= P_id_max x5. 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_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0 (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_bs (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_and (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) => P_id__minus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_l (x5::nil)) => P_id_l (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) => P_id__plus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_wb (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_ge (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_min (measure x5) | (algebra.Alg.Term algebra.F.id__sharp_ nil) => P_id__sharp_ | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_size (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6::x5::nil)) => P_id_if (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_not (measure x5) | (algebra.Alg.Term algebra.F.id_n (x7::x6::x5::nil)) => P_id_n (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_1 (x5::nil)) => P_id_1 (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_val (measure x5) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_max (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_MIN (x5:Z) := 0. Definition P_id_0_hat_1 (x5:Z) := 0. Definition P_id_SIZE (x5:Z) := 0. Definition P_id_BS (x5:Z) := 0. Definition P_id_GE (x5:Z) (x6:Z) := 0. Definition P_id_IF (x5:Z) (x6:Z) (x7:Z) := 0. Definition P_id_MAX (x5:Z) := 2* x5. Definition P_id__minus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_WB (x5:Z) := 0. Definition P_id__plus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_AND (x5:Z) (x6:Z) := 0. Definition P_id_NOT (x5:Z) := 0. Definition P_id_Marked_val (x5:Z) := 0. Lemma P_id_MIN_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MIN x6 <= P_id_MIN x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_hat_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0_hat_1 x6 <= P_id_0_hat_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SIZE_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_SIZE x6 <= P_id_SIZE x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_BS_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_BS x6 <= P_id_BS x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_GE_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_GE x6 x8 <= P_id_GE x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_IF_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_IF x6 x8 x10 <= P_id_IF x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_MAX_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MAX x6 <= P_id_MAX x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__minus__hat_1 x6 x8 <= P_id__minus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_WB_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_WB x6 <= P_id_WB x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__plus__hat_1 x6 x8 <= P_id__plus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_AND_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_AND x6 x8 <= P_id_AND x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_NOT_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_NOT x6 <= P_id_NOT x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_Marked_val x6 <= P_id_Marked_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max P_id_MIN P_id_0_hat_1 P_id_SIZE P_id_BS P_id_GE P_id_IF P_id_MAX P_id__minus__hat_1 P_id_WB P_id__plus__hat_1 P_id_AND P_id_NOT P_id_Marked_val. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_MIN (measure x5) | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0_hat_1 (measure x5) | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_SIZE (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_BS (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_GE (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6:: x5::nil)) => P_id_IF (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_MAX (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6:: x5::nil)) => P_id__minus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_WB (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6:: x5::nil)) => P_id__plus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_AND (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_NOT (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_Marked_val (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. intros ;apply P_id_MIN_monotonic;assumption. intros ;apply P_id_0_hat_1_monotonic;assumption. intros ;apply P_id_SIZE_monotonic;assumption. intros ;apply P_id_BS_monotonic;assumption. intros ;apply P_id_GE_monotonic;assumption. intros ;apply P_id_IF_monotonic;assumption. intros ;apply P_id_MAX_monotonic;assumption. intros ;apply P_id__minus__hat_1_monotonic;assumption. intros ;apply P_id_WB_monotonic;assumption. intros ;apply P_id__plus__hat_1_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_NOT_monotonic;assumption. intros ;apply P_id_Marked_val_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_7. Proof. intros x. apply well_founded_ind with (R:=fun x y => (Zwf.Zwf 0) (marked_measure x) (marked_measure y)). apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. clear x. intros x IHx. repeat ( constructor;inversion 1;subst; full_prove_ineq algebra.Alg.Term ltac:(algebra.Alg_ext.find_replacement ) algebra.EQT_ext.one_step_list_refl_trans_clos marked_measure marked_measure_star_monotonic (Zwf.Zwf 0) (interp.o_Z 0) ltac:(fun _ => R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 ) ltac:(fun _ => rewrite_and_unfold ) ltac:(fun _ => generate_pos_hyp ) ltac:(fun _ => cbv beta iota zeta delta - [Zplus Zmult Zle Zlt] in * ; try (constructor)) IHx ). Qed. End WF_DP_R_xml_0_scc_7. Definition wf_DP_R_xml_0_scc_7 := WF_DP_R_xml_0_scc_7.wf. Lemma acc_DP_R_xml_0_scc_7 : forall x y, (DP_R_xml_0_scc_7 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_7). intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). apply wf_DP_R_xml_0_scc_7. Qed. Inductive DP_R_xml_0_non_scc_8 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_8_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_8 (algebra.Alg.Term algebra.F.id_max (x2::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_8 : forall x y, (DP_R_xml_0_non_scc_8 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_7; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Inductive DP_R_xml_0_scc_9 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_9_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_scc_9 (algebra.Alg.Term algebra.F.id_min (x2::nil)) (algebra.Alg.Term algebra.F.id_min (x5::nil)) . Module WF_DP_R_xml_0_scc_9. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_0 (x5:Z) := 1 + 1* x5. Definition P_id_bs (x5:Z) := 2 + 2* x5. Definition P_id_and (x5:Z) (x6:Z) := 1 + 1* x5. Definition P_id__minus_ (x5:Z) (x6:Z) := 1* x5. Definition P_id_l (x5:Z) := 3 + 1* x5. Definition P_id__plus_ (x5:Z) (x6:Z) := 1* x5 + 1* x6. Definition P_id_wb (x5:Z) := 2 + 2* x5. Definition P_id_ge (x5:Z) (x6:Z) := 1* x6. Definition P_id_false := 0. Definition P_id_min (x5:Z) := 2* x5. Definition P_id__sharp_ := 0. Definition P_id_size (x5:Z) := 1* x5. Definition P_id_if (x5:Z) (x6:Z) (x7:Z) := 2* x6 + 1* x7. Definition P_id_not (x5:Z) := 0. Definition P_id_n (x5:Z) (x6:Z) (x7:Z) := 3 + 2* x5 + 2* x6 + 2* x7. Definition P_id_1 (x5:Z) := 1 + 1* x5. Definition P_id_val (x5:Z) := 1* x5. Definition P_id_true := 0. Definition P_id_max (x5:Z) := 1 + 3* x5. Lemma P_id_0_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0 x6 <= P_id_0 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_bs x6 <= P_id_bs x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_and x6 x8 <= P_id_and x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__minus_ x6 x8 <= P_id__minus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_l x6 <= P_id_l x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__plus_ x6 x8 <= P_id__plus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_wb x6 <= P_id_wb x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_ge x6 x8 <= P_id_ge x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_min_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_min x6 <= P_id_min x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_size x6 <= P_id_size x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_if x6 x8 x10 <= P_id_if x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_not_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_not x6 <= P_id_not x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_n x6 x8 x10 <= P_id_n x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_1 x6 <= P_id_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_val x6 <= P_id_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_max_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_max x6 <= P_id_max x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : forall x5, (0 <= x5) ->0 <= P_id_0 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_bounded : forall x5, (0 <= x5) ->0 <= P_id_bs x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_and x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__minus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_bounded : forall x5, (0 <= x5) ->0 <= P_id_l x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__plus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_bounded : forall x5, (0 <= x5) ->0 <= P_id_wb x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_ge x6 x5. 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_min_bounded : forall x5, (0 <= x5) ->0 <= P_id_min x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__sharp__bounded : 0 <= P_id__sharp_ . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_bounded : forall x5, (0 <= x5) ->0 <= P_id_size x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_if x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_not_bounded : forall x5, (0 <= x5) ->0 <= P_id_not x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_n x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_1_bounded : forall x5, (0 <= x5) ->0 <= P_id_1 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_bounded : forall x5, (0 <= x5) ->0 <= P_id_val x5. 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_max_bounded : forall x5, (0 <= x5) ->0 <= P_id_max x5. 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_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0 (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_bs (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_and (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) => P_id__minus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_l (x5::nil)) => P_id_l (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) => P_id__plus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_wb (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_ge (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_min (measure x5) | (algebra.Alg.Term algebra.F.id__sharp_ nil) => P_id__sharp_ | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_size (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6::x5::nil)) => P_id_if (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_not (measure x5) | (algebra.Alg.Term algebra.F.id_n (x7::x6::x5::nil)) => P_id_n (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_1 (x5::nil)) => P_id_1 (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_val (measure x5) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_max (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_MIN (x5:Z) := 3* x5. Definition P_id_0_hat_1 (x5:Z) := 0. Definition P_id_SIZE (x5:Z) := 0. Definition P_id_BS (x5:Z) := 0. Definition P_id_GE (x5:Z) (x6:Z) := 0. Definition P_id_IF (x5:Z) (x6:Z) (x7:Z) := 0. Definition P_id_MAX (x5:Z) := 0. Definition P_id__minus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_WB (x5:Z) := 0. Definition P_id__plus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_AND (x5:Z) (x6:Z) := 0. Definition P_id_NOT (x5:Z) := 0. Definition P_id_Marked_val (x5:Z) := 0. Lemma P_id_MIN_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MIN x6 <= P_id_MIN x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_hat_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0_hat_1 x6 <= P_id_0_hat_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SIZE_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_SIZE x6 <= P_id_SIZE x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_BS_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_BS x6 <= P_id_BS x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_GE_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_GE x6 x8 <= P_id_GE x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_IF_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_IF x6 x8 x10 <= P_id_IF x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_MAX_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MAX x6 <= P_id_MAX x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__minus__hat_1 x6 x8 <= P_id__minus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_WB_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_WB x6 <= P_id_WB x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__plus__hat_1 x6 x8 <= P_id__plus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_AND_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_AND x6 x8 <= P_id_AND x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_NOT_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_NOT x6 <= P_id_NOT x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_Marked_val x6 <= P_id_Marked_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max P_id_MIN P_id_0_hat_1 P_id_SIZE P_id_BS P_id_GE P_id_IF P_id_MAX P_id__minus__hat_1 P_id_WB P_id__plus__hat_1 P_id_AND P_id_NOT P_id_Marked_val. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_MIN (measure x5) | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0_hat_1 (measure x5) | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_SIZE (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_BS (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_GE (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6:: x5::nil)) => P_id_IF (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_MAX (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6:: x5::nil)) => P_id__minus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_WB (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6:: x5::nil)) => P_id__plus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_AND (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_NOT (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_Marked_val (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. intros ;apply P_id_MIN_monotonic;assumption. intros ;apply P_id_0_hat_1_monotonic;assumption. intros ;apply P_id_SIZE_monotonic;assumption. intros ;apply P_id_BS_monotonic;assumption. intros ;apply P_id_GE_monotonic;assumption. intros ;apply P_id_IF_monotonic;assumption. intros ;apply P_id_MAX_monotonic;assumption. intros ;apply P_id__minus__hat_1_monotonic;assumption. intros ;apply P_id_WB_monotonic;assumption. intros ;apply P_id__plus__hat_1_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_NOT_monotonic;assumption. intros ;apply P_id_Marked_val_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_9. 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_9. Definition wf_DP_R_xml_0_scc_9 := WF_DP_R_xml_0_scc_9.wf. Lemma acc_DP_R_xml_0_scc_9 : forall x y, (DP_R_xml_0_scc_9 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_9). 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_9. Qed. Inductive DP_R_xml_0_non_scc_10 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_10_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_10 (algebra.Alg.Term algebra.F.id_min (x3::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_10 : forall x y, (DP_R_xml_0_non_scc_10 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_9; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Inductive DP_R_xml_0_non_scc_11 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_11_0 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0_non_scc_11 (algebra.Alg.Term algebra.F.id_not ((algebra.Alg.Term algebra.F.id_ge (x2:: x1::nil))::nil)) (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_11 : forall x y, (DP_R_xml_0_non_scc_11 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_scc_12 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_0 : forall x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id__sharp_ nil) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x5) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id__sharp_ nil):: x1::nil)) (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) . Module WF_DP_R_xml_0_scc_12. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_0 (x5:Z) := 1 + 1* x5. Definition P_id_bs (x5:Z) := 1* x5. Definition P_id_and (x5:Z) (x6:Z) := 2* x5 + 1* x6. Definition P_id__minus_ (x5:Z) (x6:Z) := 1* x5. Definition P_id_l (x5:Z) := 1 + 1* x5. Definition P_id__plus_ (x5:Z) (x6:Z) := 1* x5 + 2* x6. Definition P_id_wb (x5:Z) := 0. Definition P_id_ge (x5:Z) (x6:Z) := 0. Definition P_id_false := 0. Definition P_id_min (x5:Z) := 2* x5. Definition P_id__sharp_ := 0. Definition P_id_size (x5:Z) := 2* x5. Definition P_id_if (x5:Z) (x6:Z) (x7:Z) := 2* x6 + 2* x7. Definition P_id_not (x5:Z) := 0. Definition P_id_n (x5:Z) (x6:Z) (x7:Z) := 2 + 3* x5 + 2* x6 + 2* x7. Definition P_id_1 (x5:Z) := 1 + 1* x5. Definition P_id_val (x5:Z) := 1* x5. Definition P_id_true := 0. Definition P_id_max (x5:Z) := 3* x5. Lemma P_id_0_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0 x6 <= P_id_0 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_bs x6 <= P_id_bs x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_and x6 x8 <= P_id_and x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__minus_ x6 x8 <= P_id__minus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_l x6 <= P_id_l x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__plus_ x6 x8 <= P_id__plus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_wb x6 <= P_id_wb x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_ge x6 x8 <= P_id_ge x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_min_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_min x6 <= P_id_min x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_size x6 <= P_id_size x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_if x6 x8 x10 <= P_id_if x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_not_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_not x6 <= P_id_not x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_n x6 x8 x10 <= P_id_n x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_1 x6 <= P_id_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_val x6 <= P_id_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_max_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_max x6 <= P_id_max x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : forall x5, (0 <= x5) ->0 <= P_id_0 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_bounded : forall x5, (0 <= x5) ->0 <= P_id_bs x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_and x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__minus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_bounded : forall x5, (0 <= x5) ->0 <= P_id_l x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__plus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_bounded : forall x5, (0 <= x5) ->0 <= P_id_wb x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_ge x6 x5. 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_min_bounded : forall x5, (0 <= x5) ->0 <= P_id_min x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__sharp__bounded : 0 <= P_id__sharp_ . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_bounded : forall x5, (0 <= x5) ->0 <= P_id_size x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_if x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_not_bounded : forall x5, (0 <= x5) ->0 <= P_id_not x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_n x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_1_bounded : forall x5, (0 <= x5) ->0 <= P_id_1 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_bounded : forall x5, (0 <= x5) ->0 <= P_id_val x5. 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_max_bounded : forall x5, (0 <= x5) ->0 <= P_id_max x5. 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_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0 (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_bs (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_and (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) => P_id__minus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_l (x5::nil)) => P_id_l (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) => P_id__plus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_wb (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_ge (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_min (measure x5) | (algebra.Alg.Term algebra.F.id__sharp_ nil) => P_id__sharp_ | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_size (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6::x5::nil)) => P_id_if (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_not (measure x5) | (algebra.Alg.Term algebra.F.id_n (x7::x6::x5::nil)) => P_id_n (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_1 (x5::nil)) => P_id_1 (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_val (measure x5) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_max (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_MIN (x5:Z) := 0. Definition P_id_0_hat_1 (x5:Z) := 0. Definition P_id_SIZE (x5:Z) := 0. Definition P_id_BS (x5:Z) := 0. Definition P_id_GE (x5:Z) (x6:Z) := 2* x6. Definition P_id_IF (x5:Z) (x6:Z) (x7:Z) := 0. Definition P_id_MAX (x5:Z) := 0. Definition P_id__minus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_WB (x5:Z) := 0. Definition P_id__plus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_AND (x5:Z) (x6:Z) := 0. Definition P_id_NOT (x5:Z) := 0. Definition P_id_Marked_val (x5:Z) := 0. Lemma P_id_MIN_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MIN x6 <= P_id_MIN x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_hat_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0_hat_1 x6 <= P_id_0_hat_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SIZE_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_SIZE x6 <= P_id_SIZE x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_BS_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_BS x6 <= P_id_BS x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_GE_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_GE x6 x8 <= P_id_GE x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_IF_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_IF x6 x8 x10 <= P_id_IF x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_MAX_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MAX x6 <= P_id_MAX x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__minus__hat_1 x6 x8 <= P_id__minus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_WB_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_WB x6 <= P_id_WB x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__plus__hat_1 x6 x8 <= P_id__plus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_AND_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_AND x6 x8 <= P_id_AND x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_NOT_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_NOT x6 <= P_id_NOT x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_Marked_val x6 <= P_id_Marked_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max P_id_MIN P_id_0_hat_1 P_id_SIZE P_id_BS P_id_GE P_id_IF P_id_MAX P_id__minus__hat_1 P_id_WB P_id__plus__hat_1 P_id_AND P_id_NOT P_id_Marked_val. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_MIN (measure x5) | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0_hat_1 (measure x5) | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_SIZE (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_BS (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_GE (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6:: x5::nil)) => P_id_IF (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_MAX (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6:: x5::nil)) => P_id__minus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_WB (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6:: x5::nil)) => P_id__plus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_AND (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_NOT (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_Marked_val (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. intros ;apply P_id_MIN_monotonic;assumption. intros ;apply P_id_0_hat_1_monotonic;assumption. intros ;apply P_id_SIZE_monotonic;assumption. intros ;apply P_id_BS_monotonic;assumption. intros ;apply P_id_GE_monotonic;assumption. intros ;apply P_id_IF_monotonic;assumption. intros ;apply P_id_MAX_monotonic;assumption. intros ;apply P_id__minus__hat_1_monotonic;assumption. intros ;apply P_id_WB_monotonic;assumption. intros ;apply P_id__plus__hat_1_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_NOT_monotonic;assumption. intros ;apply P_id_Marked_val_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_12. Proof. intros x. apply well_founded_ind with (R:=fun x y => (Zwf.Zwf 0) (marked_measure x) (marked_measure y)). apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. clear x. intros x IHx. repeat ( constructor;inversion 1;subst; full_prove_ineq algebra.Alg.Term ltac:(algebra.Alg_ext.find_replacement ) algebra.EQT_ext.one_step_list_refl_trans_clos marked_measure marked_measure_star_monotonic (Zwf.Zwf 0) (interp.o_Z 0) ltac:(fun _ => R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 ) ltac:(fun _ => rewrite_and_unfold ) ltac:(fun _ => generate_pos_hyp ) ltac:(fun _ => cbv beta iota zeta delta - [Zplus Zmult Zle Zlt] in * ; try (constructor)) IHx ). Qed. End WF_DP_R_xml_0_scc_12. Definition wf_DP_R_xml_0_scc_12 := WF_DP_R_xml_0_scc_12.wf. Lemma acc_DP_R_xml_0_scc_12 : forall x y, (DP_R_xml_0_scc_12 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_12). intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((eapply acc_DP_R_xml_0_non_scc_11; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))). apply wf_DP_R_xml_0_scc_12. Qed. Inductive DP_R_xml_0_scc_13 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_13_0 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0_scc_13 (algebra.Alg.Term algebra.F.id_ge (x2::x1::nil)) (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) (* *) | DP_R_xml_0_scc_13_1 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0_scc_13 (algebra.Alg.Term algebra.F.id_ge (x1::x2::nil)) (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) (* *) | DP_R_xml_0_scc_13_2 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0_scc_13 (algebra.Alg.Term algebra.F.id_ge (x1::x2::nil)) (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) (* *) | DP_R_xml_0_scc_13_3 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0_scc_13 (algebra.Alg.Term algebra.F.id_ge (x1::x2::nil)) (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) . Module WF_DP_R_xml_0_scc_13. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_0 (x5:Z) := 1 + 1* x5. Definition P_id_bs (x5:Z) := 3. Definition P_id_and (x5:Z) (x6:Z) := 2* x5. Definition P_id__minus_ (x5:Z) (x6:Z) := 1* x5. Definition P_id_l (x5:Z) := 1* x5. Definition P_id__plus_ (x5:Z) (x6:Z) := 1* x5 + 2* x6. Definition P_id_wb (x5:Z) := 1. Definition P_id_ge (x5:Z) (x6:Z) := 0. Definition P_id_false := 0. Definition P_id_min (x5:Z) := 3* x5. Definition P_id__sharp_ := 0. Definition P_id_size (x5:Z) := 1 + 2* x5. Definition P_id_if (x5:Z) (x6:Z) (x7:Z) := 2* x6 + 2* x7. Definition P_id_not (x5:Z) := 0. Definition P_id_n (x5:Z) (x6:Z) (x7:Z) := 3 + 1* x5 + 2* x6 + 2* x7. Definition P_id_1 (x5:Z) := 1 + 1* x5. Definition P_id_val (x5:Z) := 1* x5. Definition P_id_true := 0. Definition P_id_max (x5:Z) := 1* x5. Lemma P_id_0_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0 x6 <= P_id_0 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_bs x6 <= P_id_bs x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_and x6 x8 <= P_id_and x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__minus_ x6 x8 <= P_id__minus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_l x6 <= P_id_l x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__plus_ x6 x8 <= P_id__plus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_wb x6 <= P_id_wb x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_ge x6 x8 <= P_id_ge x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_min_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_min x6 <= P_id_min x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_size x6 <= P_id_size x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_if x6 x8 x10 <= P_id_if x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_not_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_not x6 <= P_id_not x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_n x6 x8 x10 <= P_id_n x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_1 x6 <= P_id_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_val x6 <= P_id_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_max_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_max x6 <= P_id_max x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : forall x5, (0 <= x5) ->0 <= P_id_0 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_bounded : forall x5, (0 <= x5) ->0 <= P_id_bs x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_and x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__minus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_bounded : forall x5, (0 <= x5) ->0 <= P_id_l x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__plus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_bounded : forall x5, (0 <= x5) ->0 <= P_id_wb x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_ge x6 x5. 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_min_bounded : forall x5, (0 <= x5) ->0 <= P_id_min x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__sharp__bounded : 0 <= P_id__sharp_ . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_bounded : forall x5, (0 <= x5) ->0 <= P_id_size x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_if x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_not_bounded : forall x5, (0 <= x5) ->0 <= P_id_not x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_n x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_1_bounded : forall x5, (0 <= x5) ->0 <= P_id_1 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_bounded : forall x5, (0 <= x5) ->0 <= P_id_val x5. 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_max_bounded : forall x5, (0 <= x5) ->0 <= P_id_max x5. 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_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0 (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_bs (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_and (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) => P_id__minus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_l (x5::nil)) => P_id_l (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) => P_id__plus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_wb (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_ge (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_min (measure x5) | (algebra.Alg.Term algebra.F.id__sharp_ nil) => P_id__sharp_ | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_size (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6::x5::nil)) => P_id_if (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_not (measure x5) | (algebra.Alg.Term algebra.F.id_n (x7::x6::x5::nil)) => P_id_n (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_1 (x5::nil)) => P_id_1 (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_val (measure x5) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_max (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_MIN (x5:Z) := 0. Definition P_id_0_hat_1 (x5:Z) := 0. Definition P_id_SIZE (x5:Z) := 0. Definition P_id_BS (x5:Z) := 0. Definition P_id_GE (x5:Z) (x6:Z) := 3* x5 + 3* x6. Definition P_id_IF (x5:Z) (x6:Z) (x7:Z) := 0. Definition P_id_MAX (x5:Z) := 0. Definition P_id__minus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_WB (x5:Z) := 0. Definition P_id__plus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_AND (x5:Z) (x6:Z) := 0. Definition P_id_NOT (x5:Z) := 0. Definition P_id_Marked_val (x5:Z) := 0. Lemma P_id_MIN_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MIN x6 <= P_id_MIN x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_hat_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0_hat_1 x6 <= P_id_0_hat_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SIZE_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_SIZE x6 <= P_id_SIZE x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_BS_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_BS x6 <= P_id_BS x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_GE_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_GE x6 x8 <= P_id_GE x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_IF_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_IF x6 x8 x10 <= P_id_IF x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_MAX_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MAX x6 <= P_id_MAX x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__minus__hat_1 x6 x8 <= P_id__minus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_WB_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_WB x6 <= P_id_WB x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__plus__hat_1 x6 x8 <= P_id__plus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_AND_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_AND x6 x8 <= P_id_AND x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_NOT_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_NOT x6 <= P_id_NOT x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_Marked_val x6 <= P_id_Marked_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max P_id_MIN P_id_0_hat_1 P_id_SIZE P_id_BS P_id_GE P_id_IF P_id_MAX P_id__minus__hat_1 P_id_WB P_id__plus__hat_1 P_id_AND P_id_NOT P_id_Marked_val. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_MIN (measure x5) | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0_hat_1 (measure x5) | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_SIZE (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_BS (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_GE (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6:: x5::nil)) => P_id_IF (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_MAX (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6:: x5::nil)) => P_id__minus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_WB (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6:: x5::nil)) => P_id__plus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_AND (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_NOT (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_Marked_val (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. intros ;apply P_id_MIN_monotonic;assumption. intros ;apply P_id_0_hat_1_monotonic;assumption. intros ;apply P_id_SIZE_monotonic;assumption. intros ;apply P_id_BS_monotonic;assumption. intros ;apply P_id_GE_monotonic;assumption. intros ;apply P_id_IF_monotonic;assumption. intros ;apply P_id_MAX_monotonic;assumption. intros ;apply P_id__minus__hat_1_monotonic;assumption. intros ;apply P_id_WB_monotonic;assumption. intros ;apply P_id__plus__hat_1_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_NOT_monotonic;assumption. intros ;apply P_id_Marked_val_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_13. Proof. intros x. apply well_founded_ind with (R:=fun x y => (Zwf.Zwf 0) (marked_measure x) (marked_measure y)). apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. clear x. intros x IHx. repeat ( constructor;inversion 1;subst; full_prove_ineq algebra.Alg.Term ltac:(algebra.Alg_ext.find_replacement ) algebra.EQT_ext.one_step_list_refl_trans_clos marked_measure marked_measure_star_monotonic (Zwf.Zwf 0) (interp.o_Z 0) ltac:(fun _ => R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 ) ltac:(fun _ => rewrite_and_unfold ) ltac:(fun _ => generate_pos_hyp ) ltac:(fun _ => cbv beta iota zeta delta - [Zplus Zmult Zle Zlt] in * ; try (constructor)) IHx ). Qed. End WF_DP_R_xml_0_scc_13. Definition wf_DP_R_xml_0_scc_13 := WF_DP_R_xml_0_scc_13.wf. Lemma acc_DP_R_xml_0_scc_13 : forall x y, (DP_R_xml_0_scc_13 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_13). intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((eapply acc_DP_R_xml_0_scc_12; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_11; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). apply wf_DP_R_xml_0_scc_13. Qed. Inductive DP_R_xml_0_non_scc_14 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_14_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_14 (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::(algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size (x3::nil)):: (algebra.Alg.Term algebra.F.id_size (x2::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_14 : forall x y, (DP_R_xml_0_non_scc_14 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_13; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_11; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). Qed. Inductive DP_R_xml_0_non_scc_15 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_15_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_15 (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::(algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size (x2::nil)):: (algebra.Alg.Term algebra.F.id_size (x3::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_15 : forall x y, (DP_R_xml_0_non_scc_15 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_13; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_11; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). Qed. Inductive DP_R_xml_0_non_scc_16 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_16_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_16 (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_size (x2::nil))::(algebra.Alg.Term algebra.F.id_size (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_16 : forall x y, (DP_R_xml_0_non_scc_16 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_13; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_11; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). Qed. Inductive DP_R_xml_0_non_scc_17 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_17_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_17 (algebra.Alg.Term algebra.F.id_ge ((algebra.Alg.Term algebra.F.id_min (x3::nil))::x1::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_17 : forall x y, (DP_R_xml_0_non_scc_17 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_13; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_11; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). Qed. Inductive DP_R_xml_0_non_scc_18 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_18_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_18 (algebra.Alg.Term algebra.F.id_ge (x1:: (algebra.Alg.Term algebra.F.id_max (x2::nil))::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_18 : forall x y, (DP_R_xml_0_non_scc_18 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_13; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_11; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). Qed. Inductive DP_R_xml_0_scc_19 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_19_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_scc_19 (algebra.Alg.Term algebra.F.id_bs (x3::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) (* *) | DP_R_xml_0_scc_19_1 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_scc_19 (algebra.Alg.Term algebra.F.id_bs (x2::nil)) (algebra.Alg.Term algebra.F.id_bs (x5::nil)) . Module WF_DP_R_xml_0_scc_19. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_0 (x5:Z) := 0. Definition P_id_bs (x5:Z) := 3* x5. Definition P_id_and (x5:Z) (x6:Z) := 2* x5. Definition P_id__minus_ (x5:Z) (x6:Z) := 1* x5 + 2* x6. Definition P_id_l (x5:Z) := 2 + 1* x5. Definition P_id__plus_ (x5:Z) (x6:Z) := 1* x5 + 1* x6. Definition P_id_wb (x5:Z) := 2 + 2* x5. Definition P_id_ge (x5:Z) (x6:Z) := 1. Definition P_id_false := 0. Definition P_id_min (x5:Z) := 3* x5. Definition P_id__sharp_ := 0. Definition P_id_size (x5:Z) := 1* x5. Definition P_id_if (x5:Z) (x6:Z) (x7:Z) := 1* x5 + 2* x6 + 1* x7. Definition P_id_not (x5:Z) := 1. Definition P_id_n (x5:Z) (x6:Z) (x7:Z) := 3 + 3* x5 + 3* x6 + 2* x7. Definition P_id_1 (x5:Z) := 2. Definition P_id_val (x5:Z) := 1* x5. Definition P_id_true := 1. Definition P_id_max (x5:Z) := 1* x5. Lemma P_id_0_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0 x6 <= P_id_0 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_bs x6 <= P_id_bs x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_and x6 x8 <= P_id_and x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__minus_ x6 x8 <= P_id__minus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_l x6 <= P_id_l x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__plus_ x6 x8 <= P_id__plus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_wb x6 <= P_id_wb x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_ge x6 x8 <= P_id_ge x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_min_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_min x6 <= P_id_min x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_size x6 <= P_id_size x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_if x6 x8 x10 <= P_id_if x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_not_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_not x6 <= P_id_not x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_n x6 x8 x10 <= P_id_n x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_1 x6 <= P_id_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_val x6 <= P_id_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_max_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_max x6 <= P_id_max x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : forall x5, (0 <= x5) ->0 <= P_id_0 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_bounded : forall x5, (0 <= x5) ->0 <= P_id_bs x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_and x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__minus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_bounded : forall x5, (0 <= x5) ->0 <= P_id_l x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__plus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_bounded : forall x5, (0 <= x5) ->0 <= P_id_wb x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_ge x6 x5. 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_min_bounded : forall x5, (0 <= x5) ->0 <= P_id_min x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__sharp__bounded : 0 <= P_id__sharp_ . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_bounded : forall x5, (0 <= x5) ->0 <= P_id_size x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_if x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_not_bounded : forall x5, (0 <= x5) ->0 <= P_id_not x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_n x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_1_bounded : forall x5, (0 <= x5) ->0 <= P_id_1 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_bounded : forall x5, (0 <= x5) ->0 <= P_id_val x5. 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_max_bounded : forall x5, (0 <= x5) ->0 <= P_id_max x5. 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_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0 (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_bs (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_and (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) => P_id__minus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_l (x5::nil)) => P_id_l (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) => P_id__plus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_wb (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_ge (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_min (measure x5) | (algebra.Alg.Term algebra.F.id__sharp_ nil) => P_id__sharp_ | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_size (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6::x5::nil)) => P_id_if (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_not (measure x5) | (algebra.Alg.Term algebra.F.id_n (x7::x6::x5::nil)) => P_id_n (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_1 (x5::nil)) => P_id_1 (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_val (measure x5) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_max (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_MIN (x5:Z) := 0. Definition P_id_0_hat_1 (x5:Z) := 0. Definition P_id_SIZE (x5:Z) := 0. Definition P_id_BS (x5:Z) := 2* x5. Definition P_id_GE (x5:Z) (x6:Z) := 0. Definition P_id_IF (x5:Z) (x6:Z) (x7:Z) := 0. Definition P_id_MAX (x5:Z) := 0. Definition P_id__minus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_WB (x5:Z) := 0. Definition P_id__plus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_AND (x5:Z) (x6:Z) := 0. Definition P_id_NOT (x5:Z) := 0. Definition P_id_Marked_val (x5:Z) := 0. Lemma P_id_MIN_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MIN x6 <= P_id_MIN x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_hat_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0_hat_1 x6 <= P_id_0_hat_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SIZE_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_SIZE x6 <= P_id_SIZE x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_BS_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_BS x6 <= P_id_BS x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_GE_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_GE x6 x8 <= P_id_GE x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_IF_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_IF x6 x8 x10 <= P_id_IF x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_MAX_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MAX x6 <= P_id_MAX x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__minus__hat_1 x6 x8 <= P_id__minus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_WB_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_WB x6 <= P_id_WB x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__plus__hat_1 x6 x8 <= P_id__plus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_AND_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_AND x6 x8 <= P_id_AND x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_NOT_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_NOT x6 <= P_id_NOT x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_Marked_val x6 <= P_id_Marked_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max P_id_MIN P_id_0_hat_1 P_id_SIZE P_id_BS P_id_GE P_id_IF P_id_MAX P_id__minus__hat_1 P_id_WB P_id__plus__hat_1 P_id_AND P_id_NOT P_id_Marked_val. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_MIN (measure x5) | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0_hat_1 (measure x5) | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_SIZE (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_BS (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_GE (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6:: x5::nil)) => P_id_IF (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_MAX (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6:: x5::nil)) => P_id__minus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_WB (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6:: x5::nil)) => P_id__plus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_AND (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_NOT (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_Marked_val (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. intros ;apply P_id_MIN_monotonic;assumption. intros ;apply P_id_0_hat_1_monotonic;assumption. intros ;apply P_id_SIZE_monotonic;assumption. intros ;apply P_id_BS_monotonic;assumption. intros ;apply P_id_GE_monotonic;assumption. intros ;apply P_id_IF_monotonic;assumption. intros ;apply P_id_MAX_monotonic;assumption. intros ;apply P_id__minus__hat_1_monotonic;assumption. intros ;apply P_id_WB_monotonic;assumption. intros ;apply P_id__plus__hat_1_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_NOT_monotonic;assumption. intros ;apply P_id_Marked_val_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_19. Proof. intros x. apply well_founded_ind with (R:=fun x y => (Zwf.Zwf 0) (marked_measure x) (marked_measure y)). apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. clear x. intros x IHx. repeat ( constructor;inversion 1;subst; full_prove_ineq algebra.Alg.Term ltac:(algebra.Alg_ext.find_replacement ) algebra.EQT_ext.one_step_list_refl_trans_clos marked_measure marked_measure_star_monotonic (Zwf.Zwf 0) (interp.o_Z 0) ltac:(fun _ => R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 ) ltac:(fun _ => rewrite_and_unfold ) ltac:(fun _ => generate_pos_hyp ) ltac:(fun _ => cbv beta iota zeta delta - [Zplus Zmult Zle Zlt] in * ; try (constructor)) IHx ). Qed. End WF_DP_R_xml_0_scc_19. Definition wf_DP_R_xml_0_scc_19 := WF_DP_R_xml_0_scc_19.wf. Lemma acc_DP_R_xml_0_scc_19 : forall x y, (DP_R_xml_0_scc_19 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_19). intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((eapply acc_DP_R_xml_0_non_scc_18; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_17; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_10; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_8; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_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' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))))))))). apply wf_DP_R_xml_0_scc_19. Qed. Inductive DP_R_xml_0_non_scc_20 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* <-(1(x_),1(y_)),0(-(x_,y_))> *) | DP_R_xml_0_non_scc_20_0 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0_non_scc_20 (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__minus_ (x1:: x2::nil))::nil)) (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_20 : forall x y, (DP_R_xml_0_non_scc_20 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_21 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* <-(0(x_),0(y_)),0(-(x_,y_))> *) | DP_R_xml_0_non_scc_21_0 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0_non_scc_21 (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__minus_ (x1:: x2::nil))::nil)) (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_21 : forall x y, (DP_R_xml_0_non_scc_21 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_scc_22 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* <-(0(x_),1(y_)),-(-(x_,y_),1(#))> *) | DP_R_xml_0_scc_22_0 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0_scc_22 (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id__minus_ (x1:: x2::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) (* <-(0(x_),1(y_)),-(x_,y_)> *) | DP_R_xml_0_scc_22_1 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0_scc_22 (algebra.Alg.Term algebra.F.id__minus_ (x1:: x2::nil)) (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) (* <-(0(x_),0(y_)),-(x_,y_)> *) | DP_R_xml_0_scc_22_2 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0_scc_22 (algebra.Alg.Term algebra.F.id__minus_ (x1:: x2::nil)) (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) (* <-(1(x_),0(y_)),-(x_,y_)> *) | DP_R_xml_0_scc_22_3 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0_scc_22 (algebra.Alg.Term algebra.F.id__minus_ (x1:: x2::nil)) (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) (* <-(1(x_),1(y_)),-(x_,y_)> *) | DP_R_xml_0_scc_22_4 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0_scc_22 (algebra.Alg.Term algebra.F.id__minus_ (x1:: x2::nil)) (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) . Module WF_DP_R_xml_0_scc_22. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_0 (x5:Z) := 2 + 1* x5. Definition P_id_bs (x5:Z) := 2* x5. Definition P_id_and (x5:Z) (x6:Z) := 1 + 2* x5. Definition P_id__minus_ (x5:Z) (x6:Z) := 1* x5. Definition P_id_l (x5:Z) := 2 + 1* x5. Definition P_id__plus_ (x5:Z) (x6:Z) := 1* x5 + 1* x6. Definition P_id_wb (x5:Z) := 2. Definition P_id_ge (x5:Z) (x6:Z) := 0. Definition P_id_false := 0. Definition P_id_min (x5:Z) := 3* x5. Definition P_id__sharp_ := 0. Definition P_id_size (x5:Z) := 1 + 2* x5. Definition P_id_if (x5:Z) (x6:Z) (x7:Z) := 2* x6 + 2* x7. Definition P_id_not (x5:Z) := 0. Definition P_id_n (x5:Z) (x6:Z) (x7:Z) := 2 + 1* x5 + 2* x6 + 2* x7. Definition P_id_1 (x5:Z) := 2 + 1* x5. Definition P_id_val (x5:Z) := 1* x5. Definition P_id_true := 0. Definition P_id_max (x5:Z) := 1* x5. Lemma P_id_0_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0 x6 <= P_id_0 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_bs x6 <= P_id_bs x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_and x6 x8 <= P_id_and x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__minus_ x6 x8 <= P_id__minus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_l x6 <= P_id_l x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__plus_ x6 x8 <= P_id__plus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_wb x6 <= P_id_wb x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_ge x6 x8 <= P_id_ge x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_min_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_min x6 <= P_id_min x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_size x6 <= P_id_size x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_if x6 x8 x10 <= P_id_if x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_not_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_not x6 <= P_id_not x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_n x6 x8 x10 <= P_id_n x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_1 x6 <= P_id_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_val x6 <= P_id_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_max_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_max x6 <= P_id_max x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : forall x5, (0 <= x5) ->0 <= P_id_0 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_bounded : forall x5, (0 <= x5) ->0 <= P_id_bs x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_and x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__minus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_bounded : forall x5, (0 <= x5) ->0 <= P_id_l x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__plus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_bounded : forall x5, (0 <= x5) ->0 <= P_id_wb x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_ge x6 x5. 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_min_bounded : forall x5, (0 <= x5) ->0 <= P_id_min x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__sharp__bounded : 0 <= P_id__sharp_ . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_bounded : forall x5, (0 <= x5) ->0 <= P_id_size x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_if x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_not_bounded : forall x5, (0 <= x5) ->0 <= P_id_not x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_n x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_1_bounded : forall x5, (0 <= x5) ->0 <= P_id_1 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_bounded : forall x5, (0 <= x5) ->0 <= P_id_val x5. 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_max_bounded : forall x5, (0 <= x5) ->0 <= P_id_max x5. 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_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0 (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_bs (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_and (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) => P_id__minus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_l (x5::nil)) => P_id_l (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) => P_id__plus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_wb (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_ge (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_min (measure x5) | (algebra.Alg.Term algebra.F.id__sharp_ nil) => P_id__sharp_ | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_size (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6::x5::nil)) => P_id_if (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_not (measure x5) | (algebra.Alg.Term algebra.F.id_n (x7::x6::x5::nil)) => P_id_n (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_1 (x5::nil)) => P_id_1 (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_val (measure x5) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_max (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_MIN (x5:Z) := 0. Definition P_id_0_hat_1 (x5:Z) := 0. Definition P_id_SIZE (x5:Z) := 0. Definition P_id_BS (x5:Z) := 0. Definition P_id_GE (x5:Z) (x6:Z) := 0. Definition P_id_IF (x5:Z) (x6:Z) (x7:Z) := 0. Definition P_id_MAX (x5:Z) := 0. Definition P_id__minus__hat_1 (x5:Z) (x6:Z) := 2* x5. Definition P_id_WB (x5:Z) := 0. Definition P_id__plus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_AND (x5:Z) (x6:Z) := 0. Definition P_id_NOT (x5:Z) := 0. Definition P_id_Marked_val (x5:Z) := 0. Lemma P_id_MIN_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MIN x6 <= P_id_MIN x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_hat_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0_hat_1 x6 <= P_id_0_hat_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SIZE_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_SIZE x6 <= P_id_SIZE x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_BS_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_BS x6 <= P_id_BS x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_GE_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_GE x6 x8 <= P_id_GE x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_IF_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_IF x6 x8 x10 <= P_id_IF x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_MAX_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MAX x6 <= P_id_MAX x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__minus__hat_1 x6 x8 <= P_id__minus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_WB_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_WB x6 <= P_id_WB x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__plus__hat_1 x6 x8 <= P_id__plus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_AND_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_AND x6 x8 <= P_id_AND x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_NOT_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_NOT x6 <= P_id_NOT x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_Marked_val x6 <= P_id_Marked_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max P_id_MIN P_id_0_hat_1 P_id_SIZE P_id_BS P_id_GE P_id_IF P_id_MAX P_id__minus__hat_1 P_id_WB P_id__plus__hat_1 P_id_AND P_id_NOT P_id_Marked_val. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_MIN (measure x5) | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0_hat_1 (measure x5) | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_SIZE (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_BS (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_GE (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6:: x5::nil)) => P_id_IF (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_MAX (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6:: x5::nil)) => P_id__minus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_WB (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6:: x5::nil)) => P_id__plus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_AND (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_NOT (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_Marked_val (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. intros ;apply P_id_MIN_monotonic;assumption. intros ;apply P_id_0_hat_1_monotonic;assumption. intros ;apply P_id_SIZE_monotonic;assumption. intros ;apply P_id_BS_monotonic;assumption. intros ;apply P_id_GE_monotonic;assumption. intros ;apply P_id_IF_monotonic;assumption. intros ;apply P_id_MAX_monotonic;assumption. intros ;apply P_id__minus__hat_1_monotonic;assumption. intros ;apply P_id_WB_monotonic;assumption. intros ;apply P_id__plus__hat_1_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_NOT_monotonic;assumption. intros ;apply P_id_Marked_val_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_22. 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_22. Definition wf_DP_R_xml_0_scc_22 := WF_DP_R_xml_0_scc_22.wf. Lemma acc_DP_R_xml_0_scc_22 : forall x y, (DP_R_xml_0_scc_22 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_22). 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_21; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_20; 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_22. Qed. Inductive DP_R_xml_0_non_scc_23 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_23_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_23 (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size (x3::nil))::(algebra.Alg.Term algebra.F.id_size (x2::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_23 : forall x y, (DP_R_xml_0_non_scc_23 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_22; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_21; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_20; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). Qed. Inductive DP_R_xml_0_non_scc_24 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_24_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_24 (algebra.Alg.Term algebra.F.id__minus_ ((algebra.Alg.Term algebra.F.id_size (x2::nil))::(algebra.Alg.Term algebra.F.id_size (x3::nil))::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_24 : forall x y, (DP_R_xml_0_non_scc_24 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_22; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_21; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_20; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). Qed. Inductive DP_R_xml_0_non_scc_25 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* <+(1(x_),1(y_)),0(+(+(x_,y_),1(#)))> *) | DP_R_xml_0_non_scc_25_0 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0_non_scc_25 (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ (x1:: x2::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_25 : forall x y, (DP_R_xml_0_non_scc_25 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_26 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* <+(0(x_),0(y_)),0(+(x_,y_))> *) | DP_R_xml_0_non_scc_26_0 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0_non_scc_26 (algebra.Alg.Term algebra.F.id_0 ((algebra.Alg.Term algebra.F.id__plus_ (x1:: x2::nil))::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_26 : forall x y, (DP_R_xml_0_non_scc_26 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_scc_27 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* <+(0(x_),1(y_)),+(x_,y_)> *) | DP_R_xml_0_scc_27_0 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0_scc_27 (algebra.Alg.Term algebra.F.id__plus_ (x1:: x2::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(0(x_),0(y_)),+(x_,y_)> *) | DP_R_xml_0_scc_27_1 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0_scc_27 (algebra.Alg.Term algebra.F.id__plus_ (x1:: x2::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(1(x_),0(y_)),+(x_,y_)> *) | DP_R_xml_0_scc_27_2 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0_scc_27 (algebra.Alg.Term algebra.F.id__plus_ (x1:: x2::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(1(x_),1(y_)),+(+(x_,y_),1(#))> *) | DP_R_xml_0_scc_27_3 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0_scc_27 (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ (x1:: x2::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(1(x_),1(y_)),+(x_,y_)> *) | DP_R_xml_0_scc_27_4 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0_scc_27 (algebra.Alg.Term algebra.F.id__plus_ (x1:: x2::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(x_,+(y_,z_)),+(+(x_,y_),z_)> *) | DP_R_xml_0_scc_27_5 : forall x2 x6 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id__plus_ (x2::x3::nil)) x5) -> DP_R_xml_0_scc_27 (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ (x1:: x2::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(x_,+(y_,z_)),+(x_,y_)> *) | DP_R_xml_0_scc_27_6 : forall x2 x6 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id__plus_ (x2::x3::nil)) x5) -> DP_R_xml_0_scc_27 (algebra.Alg.Term algebra.F.id__plus_ (x1:: x2::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) . Module WF_DP_R_xml_0_scc_27. Inductive DP_R_xml_0_scc_27_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* <+(1(x_),1(y_)),+(+(x_,y_),1(#))> *) | DP_R_xml_0_scc_27_large_0 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0_scc_27_large (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ (x1::x2::nil))::(algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(x_,+(y_,z_)),+(+(x_,y_),z_)> *) | DP_R_xml_0_scc_27_large_1 : forall x2 x6 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id__plus_ (x2::x3::nil)) x5) -> DP_R_xml_0_scc_27_large (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ (x1::x2::nil))::x3::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(x_,+(y_,z_)),+(x_,y_)> *) | DP_R_xml_0_scc_27_large_2 : forall x2 x6 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id__plus_ (x2::x3::nil)) x5) -> DP_R_xml_0_scc_27_large (algebra.Alg.Term algebra.F.id__plus_ (x1:: x2::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) . Inductive DP_R_xml_0_scc_27_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* <+(0(x_),1(y_)),+(x_,y_)> *) | DP_R_xml_0_scc_27_strict_0 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0_scc_27_strict (algebra.Alg.Term algebra.F.id__plus_ (x1:: x2::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(0(x_),0(y_)),+(x_,y_)> *) | DP_R_xml_0_scc_27_strict_1 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0_scc_27_strict (algebra.Alg.Term algebra.F.id__plus_ (x1:: x2::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(1(x_),0(y_)),+(x_,y_)> *) | DP_R_xml_0_scc_27_strict_2 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_0 (x2::nil)) x5) -> DP_R_xml_0_scc_27_strict (algebra.Alg.Term algebra.F.id__plus_ (x1:: x2::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(1(x_),1(y_)),+(x_,y_)> *) | DP_R_xml_0_scc_27_strict_3 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0_scc_27_strict (algebra.Alg.Term algebra.F.id__plus_ (x1:: x2::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) . Module WF_DP_R_xml_0_scc_27_large. Inductive DP_R_xml_0_scc_27_large_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* <+(1(x_),1(y_)),+(+(x_,y_),1(#))> *) | DP_R_xml_0_scc_27_large_scc_1_0 : forall x2 x6 x1 x5, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x1::nil)) x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_1 (x2::nil)) x5) -> DP_R_xml_0_scc_27_large_scc_1 (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ (x1::x2::nil)):: (algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) . Module WF_DP_R_xml_0_scc_27_large_scc_1. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_0 (x5:Z) := 2 + 1* x5. Definition P_id_bs (x5:Z) := 3. Definition P_id_and (x5:Z) (x6:Z) := 2* x5. Definition P_id__minus_ (x5:Z) (x6:Z) := 1* x5. Definition P_id_l (x5:Z) := 1 + 1* x5. Definition P_id__plus_ (x5:Z) (x6:Z) := 1* x5 + 2* x6. Definition P_id_wb (x5:Z) := 1. Definition P_id_ge (x5:Z) (x6:Z) := 0. Definition P_id_false := 0. Definition P_id_min (x5:Z) := 3* x5. Definition P_id__sharp_ := 0. Definition P_id_size (x5:Z) := 2* x5. Definition P_id_if (x5:Z) (x6:Z) (x7:Z) := 2* x6 + 2* x7. Definition P_id_not (x5:Z) := 0. Definition P_id_n (x5:Z) (x6:Z) (x7:Z) := 3 + 2* x5 + 2* x6 + 2* x7. Definition P_id_1 (x5:Z) := 2 + 1* x5. Definition P_id_val (x5:Z) := 1* x5. Definition P_id_true := 0. Definition P_id_max (x5:Z) := 1* x5. Lemma P_id_0_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0 x6 <= P_id_0 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_bs x6 <= P_id_bs x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_and x6 x8 <= P_id_and x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__minus_ x6 x8 <= P_id__minus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_l x6 <= P_id_l x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__plus_ x6 x8 <= P_id__plus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_wb x6 <= P_id_wb x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_ge x6 x8 <= P_id_ge x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_min_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_min x6 <= P_id_min x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_size x6 <= P_id_size x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_if x6 x8 x10 <= P_id_if x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_not_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_not x6 <= P_id_not x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_n x6 x8 x10 <= P_id_n x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_1 x6 <= P_id_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_val x6 <= P_id_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_max_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_max x6 <= P_id_max x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : forall x5, (0 <= x5) ->0 <= P_id_0 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_bounded : forall x5, (0 <= x5) ->0 <= P_id_bs x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_and x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__minus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_bounded : forall x5, (0 <= x5) ->0 <= P_id_l x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__plus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_bounded : forall x5, (0 <= x5) ->0 <= P_id_wb x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_ge x6 x5. 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_min_bounded : forall x5, (0 <= x5) ->0 <= P_id_min x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__sharp__bounded : 0 <= P_id__sharp_ . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_bounded : forall x5, (0 <= x5) ->0 <= P_id_size x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_if x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_not_bounded : forall x5, (0 <= x5) ->0 <= P_id_not x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_n x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_1_bounded : forall x5, (0 <= x5) ->0 <= P_id_1 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_bounded : forall x5, (0 <= x5) ->0 <= P_id_val x5. 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_max_bounded : forall x5, (0 <= x5) ->0 <= P_id_max x5. 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_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0 (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_bs (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_and (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) => P_id__minus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_l (x5::nil)) => P_id_l (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) => P_id__plus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_wb (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_ge (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_min (measure x5) | (algebra.Alg.Term algebra.F.id__sharp_ nil) => P_id__sharp_ | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_size (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6::x5::nil)) => P_id_if (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_not (measure x5) | (algebra.Alg.Term algebra.F.id_n (x7::x6::x5::nil)) => P_id_n (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_1 (x5::nil)) => P_id_1 (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_val (measure x5) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_max (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_MIN (x5:Z) := 0. Definition P_id_0_hat_1 (x5:Z) := 0. Definition P_id_SIZE (x5:Z) := 0. Definition P_id_BS (x5:Z) := 0. Definition P_id_GE (x5:Z) (x6:Z) := 0. Definition P_id_IF (x5:Z) (x6:Z) (x7:Z) := 0. Definition P_id_MAX (x5:Z) := 0. Definition P_id__minus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_WB (x5:Z) := 0. Definition P_id__plus__hat_1 (x5:Z) (x6:Z) := 1* x5 + 2* x6. Definition P_id_AND (x5:Z) (x6:Z) := 0. Definition P_id_NOT (x5:Z) := 0. Definition P_id_Marked_val (x5:Z) := 0. Lemma P_id_MIN_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MIN x6 <= P_id_MIN x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_hat_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0_hat_1 x6 <= P_id_0_hat_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SIZE_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_SIZE x6 <= P_id_SIZE x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_BS_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_BS x6 <= P_id_BS x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_GE_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_GE x6 x8 <= P_id_GE x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_IF_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_IF x6 x8 x10 <= P_id_IF x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_MAX_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MAX x6 <= P_id_MAX x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__minus__hat_1 x6 x8 <= P_id__minus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_WB_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_WB x6 <= P_id_WB x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__plus__hat_1 x6 x8 <= P_id__plus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_AND_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_AND x6 x8 <= P_id_AND x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_NOT_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_NOT x6 <= P_id_NOT x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_Marked_val x6 <= P_id_Marked_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max P_id_MIN P_id_0_hat_1 P_id_SIZE P_id_BS P_id_GE P_id_IF P_id_MAX P_id__minus__hat_1 P_id_WB P_id__plus__hat_1 P_id_AND P_id_NOT P_id_Marked_val. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_MIN (measure x5) | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0_hat_1 (measure x5) | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_SIZE (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_BS (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6:: x5::nil)) => P_id_GE (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6:: x5::nil)) => P_id_IF (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_MAX (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6:: x5::nil)) => P_id__minus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_WB (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6:: x5::nil)) => P_id__plus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6:: x5::nil)) => P_id_AND (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_NOT (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_Marked_val (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. intros ;apply P_id_MIN_monotonic;assumption. intros ;apply P_id_0_hat_1_monotonic;assumption. intros ;apply P_id_SIZE_monotonic;assumption. intros ;apply P_id_BS_monotonic;assumption. intros ;apply P_id_GE_monotonic;assumption. intros ;apply P_id_IF_monotonic;assumption. intros ;apply P_id_MAX_monotonic;assumption. intros ;apply P_id__minus__hat_1_monotonic;assumption. intros ;apply P_id_WB_monotonic;assumption. intros ;apply P_id__plus__hat_1_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_NOT_monotonic;assumption. intros ;apply P_id_Marked_val_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Lemma wf : well_founded WF_DP_R_xml_0_scc_27_large.DP_R_xml_0_scc_27_large_scc_1. Proof. intros x. apply well_founded_ind with (R:=fun x y => (Zwf.Zwf 0) (marked_measure x) (marked_measure y)). apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. clear x. intros x IHx. repeat ( constructor;inversion 1;subst; full_prove_ineq algebra.Alg.Term ltac:(algebra.Alg_ext.find_replacement ) algebra.EQT_ext.one_step_list_refl_trans_clos marked_measure marked_measure_star_monotonic (Zwf.Zwf 0) (interp.o_Z 0) ltac:(fun _ => R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 ) ltac:(fun _ => rewrite_and_unfold ) ltac:(fun _ => generate_pos_hyp ) ltac:(fun _ => cbv beta iota zeta delta - [Zplus Zmult Zle Zlt] in * ; try (constructor)) IHx ). Qed. End WF_DP_R_xml_0_scc_27_large_scc_1. Definition wf_DP_R_xml_0_scc_27_large_scc_1 := WF_DP_R_xml_0_scc_27_large_scc_1.wf. Lemma acc_DP_R_xml_0_scc_27_large_scc_1 : forall x y, (DP_R_xml_0_scc_27_large_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_27.DP_R_xml_0_scc_27_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_27_large_scc_1). intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). apply wf_DP_R_xml_0_scc_27_large_scc_1. Qed. Inductive DP_R_xml_0_scc_27_large_scc_2 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* <+(x_,+(y_,z_)),+(x_,y_)> *) | DP_R_xml_0_scc_27_large_scc_2_0 : forall x2 x6 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id__plus_ (x2::x3::nil)) x5) -> DP_R_xml_0_scc_27_large_scc_2 (algebra.Alg.Term algebra.F.id__plus_ (x1::x2::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) (* <+(x_,+(y_,z_)),+(+(x_,y_),z_)> *) | DP_R_xml_0_scc_27_large_scc_2_1 : forall x2 x6 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x1 x6) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id__plus_ (x2::x3::nil)) x5) -> DP_R_xml_0_scc_27_large_scc_2 (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ (x1::x2::nil)):: x3::nil)) (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) . Module WF_DP_R_xml_0_scc_27_large_scc_2. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_0 (x5:Z) := 1. Definition P_id_bs (x5:Z) := 1. Definition P_id_and (x5:Z) (x6:Z) := 1* x5. Definition P_id__minus_ (x5:Z) (x6:Z) := 1 + 1* x5. Definition P_id_l (x5:Z) := 1* x5. Definition P_id__plus_ (x5:Z) (x6:Z) := 1 + 1* x5 + 2* x6. Definition P_id_wb (x5:Z) := 1 + 2* x5. Definition P_id_ge (x5:Z) (x6:Z) := 0. Definition P_id_false := 0. Definition P_id_min (x5:Z) := 1* x5. Definition P_id__sharp_ := 0. Definition P_id_size (x5:Z) := 3* x5. Definition P_id_if (x5:Z) (x6:Z) (x7:Z) := 1* x6 + 1* x7. Definition P_id_not (x5:Z) := 0. Definition P_id_n (x5:Z) (x6:Z) (x7:Z) := 1 + 2* x5 + 2* x6 + 2* x7. Definition P_id_1 (x5:Z) := 0. Definition P_id_val (x5:Z) := 1* x5. Definition P_id_true := 0. Definition P_id_max (x5:Z) := 1* x5. Lemma P_id_0_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0 x6 <= P_id_0 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_bs x6 <= P_id_bs x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_and x6 x8 <= P_id_and x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__minus_ x6 x8 <= P_id__minus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_l x6 <= P_id_l x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__plus_ x6 x8 <= P_id__plus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_wb x6 <= P_id_wb x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_ge x6 x8 <= P_id_ge x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_min_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_min x6 <= P_id_min x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_size x6 <= P_id_size x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_if x6 x8 x10 <= P_id_if x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_not_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_not x6 <= P_id_not x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_n x6 x8 x10 <= P_id_n x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_1 x6 <= P_id_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_val x6 <= P_id_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_max_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_max x6 <= P_id_max x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : forall x5, (0 <= x5) ->0 <= P_id_0 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_bounded : forall x5, (0 <= x5) ->0 <= P_id_bs x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_and x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__minus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_bounded : forall x5, (0 <= x5) ->0 <= P_id_l x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__plus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_bounded : forall x5, (0 <= x5) ->0 <= P_id_wb x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_ge x6 x5. 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_min_bounded : forall x5, (0 <= x5) ->0 <= P_id_min x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__sharp__bounded : 0 <= P_id__sharp_ . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_bounded : forall x5, (0 <= x5) ->0 <= P_id_size x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_if x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_not_bounded : forall x5, (0 <= x5) ->0 <= P_id_not x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_n x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_1_bounded : forall x5, (0 <= x5) ->0 <= P_id_1 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_bounded : forall x5, (0 <= x5) ->0 <= P_id_val x5. 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_max_bounded : forall x5, (0 <= x5) ->0 <= P_id_max x5. 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_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0 (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_bs (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_and (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) => P_id__minus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_l (x5::nil)) => P_id_l (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) => P_id__plus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_wb (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_ge (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_min (measure x5) | (algebra.Alg.Term algebra.F.id__sharp_ nil) => P_id__sharp_ | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_size (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6::x5::nil)) => P_id_if (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_not (measure x5) | (algebra.Alg.Term algebra.F.id_n (x7::x6::x5::nil)) => P_id_n (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_1 (x5::nil)) => P_id_1 (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_val (measure x5) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_max (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_MIN (x5:Z) := 0. Definition P_id_0_hat_1 (x5:Z) := 0. Definition P_id_SIZE (x5:Z) := 0. Definition P_id_BS (x5:Z) := 0. Definition P_id_GE (x5:Z) (x6:Z) := 0. Definition P_id_IF (x5:Z) (x6:Z) (x7:Z) := 0. Definition P_id_MAX (x5:Z) := 0. Definition P_id__minus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_WB (x5:Z) := 0. Definition P_id__plus__hat_1 (x5:Z) (x6:Z) := 1* x6. Definition P_id_AND (x5:Z) (x6:Z) := 0. Definition P_id_NOT (x5:Z) := 0. Definition P_id_Marked_val (x5:Z) := 0. Lemma P_id_MIN_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MIN x6 <= P_id_MIN x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_hat_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0_hat_1 x6 <= P_id_0_hat_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SIZE_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_SIZE x6 <= P_id_SIZE x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_BS_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_BS x6 <= P_id_BS x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_GE_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_GE x6 x8 <= P_id_GE x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_IF_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_IF x6 x8 x10 <= P_id_IF x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_MAX_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MAX x6 <= P_id_MAX x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__minus__hat_1 x6 x8 <= P_id__minus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_WB_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_WB x6 <= P_id_WB x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__plus__hat_1 x6 x8 <= P_id__plus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_AND_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_AND x6 x8 <= P_id_AND x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_NOT_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_NOT x6 <= P_id_NOT x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_Marked_val x6 <= P_id_Marked_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max P_id_MIN P_id_0_hat_1 P_id_SIZE P_id_BS P_id_GE P_id_IF P_id_MAX P_id__minus__hat_1 P_id_WB P_id__plus__hat_1 P_id_AND P_id_NOT P_id_Marked_val. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_MIN (measure x5) | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0_hat_1 (measure x5) | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_SIZE (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_BS (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6:: x5::nil)) => P_id_GE (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6:: x5::nil)) => P_id_IF (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_MAX (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6:: x5::nil)) => P_id__minus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_WB (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6:: x5::nil)) => P_id__plus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6:: x5::nil)) => P_id_AND (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_NOT (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_Marked_val (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. intros ;apply P_id_MIN_monotonic;assumption. intros ;apply P_id_0_hat_1_monotonic;assumption. intros ;apply P_id_SIZE_monotonic;assumption. intros ;apply P_id_BS_monotonic;assumption. intros ;apply P_id_GE_monotonic;assumption. intros ;apply P_id_IF_monotonic;assumption. intros ;apply P_id_MAX_monotonic;assumption. intros ;apply P_id__minus__hat_1_monotonic;assumption. intros ;apply P_id_WB_monotonic;assumption. intros ;apply P_id__plus__hat_1_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_NOT_monotonic;assumption. intros ;apply P_id_Marked_val_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Lemma wf : well_founded WF_DP_R_xml_0_scc_27_large.DP_R_xml_0_scc_27_large_scc_2. Proof. intros x. apply well_founded_ind with (R:=fun x y => (Zwf.Zwf 0) (marked_measure x) (marked_measure y)). apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. clear x. intros x IHx. repeat ( constructor;inversion 1;subst; full_prove_ineq algebra.Alg.Term ltac:(algebra.Alg_ext.find_replacement ) algebra.EQT_ext.one_step_list_refl_trans_clos marked_measure marked_measure_star_monotonic (Zwf.Zwf 0) (interp.o_Z 0) ltac:(fun _ => R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 ) ltac:(fun _ => rewrite_and_unfold ) ltac:(fun _ => generate_pos_hyp ) ltac:(fun _ => cbv beta iota zeta delta - [Zplus Zmult Zle Zlt] in * ; try (constructor)) IHx ). Qed. End WF_DP_R_xml_0_scc_27_large_scc_2. Definition wf_DP_R_xml_0_scc_27_large_scc_2 := WF_DP_R_xml_0_scc_27_large_scc_2.wf. Lemma acc_DP_R_xml_0_scc_27_large_scc_2 : forall x y, (DP_R_xml_0_scc_27_large_scc_2 x y) -> Acc WF_DP_R_xml_0_scc_27.DP_R_xml_0_scc_27_large x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_27_large_scc_2). intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((eapply acc_DP_R_xml_0_scc_27_large_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))). apply wf_DP_R_xml_0_scc_27_large_scc_2. Qed. Lemma wf : well_founded WF_DP_R_xml_0_scc_27.DP_R_xml_0_scc_27_large. Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_27_large_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_27_large_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_27_large_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||(fail)))). Qed. End WF_DP_R_xml_0_scc_27_large. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_0 (x5:Z) := 1 + 1* x5. Definition P_id_bs (x5:Z) := 0. Definition P_id_and (x5:Z) (x6:Z) := 2* x5. Definition P_id__minus_ (x5:Z) (x6:Z) := 1* x5. Definition P_id_l (x5:Z) := 3 + 1* x5. Definition P_id__plus_ (x5:Z) (x6:Z) := 1* x5 + 1* x6. Definition P_id_wb (x5:Z) := 0. Definition P_id_ge (x5:Z) (x6:Z) := 0. Definition P_id_false := 0. Definition P_id_min (x5:Z) := 2* x5. Definition P_id__sharp_ := 0. Definition P_id_size (x5:Z) := 2 + 2* x5. Definition P_id_if (x5:Z) (x6:Z) (x7:Z) := 2* x6 + 2* x7. Definition P_id_not (x5:Z) := 0. Definition P_id_n (x5:Z) (x6:Z) (x7:Z) := 3 + 1* x5 + 1* x6 + 2* x7. Definition P_id_1 (x5:Z) := 1 + 1* x5. Definition P_id_val (x5:Z) := 1* x5. Definition P_id_true := 0. Definition P_id_max (x5:Z) := 2* x5. Lemma P_id_0_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0 x6 <= P_id_0 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_bs x6 <= P_id_bs x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_and x6 x8 <= P_id_and x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__minus_ x6 x8 <= P_id__minus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_l x6 <= P_id_l x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__plus_ x6 x8 <= P_id__plus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_wb x6 <= P_id_wb x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_ge x6 x8 <= P_id_ge x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_min_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_min x6 <= P_id_min x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_size x6 <= P_id_size x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_if x6 x8 x10 <= P_id_if x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_not_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_not x6 <= P_id_not x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_n x6 x8 x10 <= P_id_n x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_1 x6 <= P_id_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_val x6 <= P_id_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_max_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_max x6 <= P_id_max x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : forall x5, (0 <= x5) ->0 <= P_id_0 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_bounded : forall x5, (0 <= x5) ->0 <= P_id_bs x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_and x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__minus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_bounded : forall x5, (0 <= x5) ->0 <= P_id_l x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__plus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_bounded : forall x5, (0 <= x5) ->0 <= P_id_wb x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_ge x6 x5. 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_min_bounded : forall x5, (0 <= x5) ->0 <= P_id_min x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__sharp__bounded : 0 <= P_id__sharp_ . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_bounded : forall x5, (0 <= x5) ->0 <= P_id_size x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_if x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_not_bounded : forall x5, (0 <= x5) ->0 <= P_id_not x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_n x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_1_bounded : forall x5, (0 <= x5) ->0 <= P_id_1 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_bounded : forall x5, (0 <= x5) ->0 <= P_id_val x5. 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_max_bounded : forall x5, (0 <= x5) ->0 <= P_id_max x5. 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_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0 (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_bs (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_and (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) => P_id__minus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_l (x5::nil)) => P_id_l (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) => P_id__plus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_wb (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_ge (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_min (measure x5) | (algebra.Alg.Term algebra.F.id__sharp_ nil) => P_id__sharp_ | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_size (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6::x5::nil)) => P_id_if (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_not (measure x5) | (algebra.Alg.Term algebra.F.id_n (x7::x6::x5::nil)) => P_id_n (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_1 (x5::nil)) => P_id_1 (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_val (measure x5) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_max (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_MIN (x5:Z) := 0. Definition P_id_0_hat_1 (x5:Z) := 0. Definition P_id_SIZE (x5:Z) := 0. Definition P_id_BS (x5:Z) := 0. Definition P_id_GE (x5:Z) (x6:Z) := 0. Definition P_id_IF (x5:Z) (x6:Z) (x7:Z) := 0. Definition P_id_MAX (x5:Z) := 0. Definition P_id__minus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_WB (x5:Z) := 0. Definition P_id__plus__hat_1 (x5:Z) (x6:Z) := 2* x6. Definition P_id_AND (x5:Z) (x6:Z) := 0. Definition P_id_NOT (x5:Z) := 0. Definition P_id_Marked_val (x5:Z) := 0. Lemma P_id_MIN_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MIN x6 <= P_id_MIN x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_hat_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0_hat_1 x6 <= P_id_0_hat_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SIZE_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_SIZE x6 <= P_id_SIZE x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_BS_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_BS x6 <= P_id_BS x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_GE_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_GE x6 x8 <= P_id_GE x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_IF_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_IF x6 x8 x10 <= P_id_IF x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_MAX_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MAX x6 <= P_id_MAX x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__minus__hat_1 x6 x8 <= P_id__minus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_WB_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_WB x6 <= P_id_WB x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__plus__hat_1 x6 x8 <= P_id__plus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_AND_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_AND x6 x8 <= P_id_AND x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_NOT_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_NOT x6 <= P_id_NOT x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_Marked_val x6 <= P_id_Marked_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max P_id_MIN P_id_0_hat_1 P_id_SIZE P_id_BS P_id_GE P_id_IF P_id_MAX P_id__minus__hat_1 P_id_WB P_id__plus__hat_1 P_id_AND P_id_NOT P_id_Marked_val. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_MIN (measure x5) | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0_hat_1 (measure x5) | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_SIZE (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_BS (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_GE (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6:: x5::nil)) => P_id_IF (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_MAX (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6:: x5::nil)) => P_id__minus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_WB (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6:: x5::nil)) => P_id__plus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_AND (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_NOT (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_Marked_val (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. intros ;apply P_id_MIN_monotonic;assumption. intros ;apply P_id_0_hat_1_monotonic;assumption. intros ;apply P_id_SIZE_monotonic;assumption. intros ;apply P_id_BS_monotonic;assumption. intros ;apply P_id_GE_monotonic;assumption. intros ;apply P_id_IF_monotonic;assumption. intros ;apply P_id_MAX_monotonic;assumption. intros ;apply P_id__minus__hat_1_monotonic;assumption. intros ;apply P_id_WB_monotonic;assumption. intros ;apply P_id__plus__hat_1_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_NOT_monotonic;assumption. intros ;apply P_id_Marked_val_monotonic;assumption. Qed. Ltac rewrite_and_unfold := do 2 (rewrite marked_measure_equation); repeat ( match goal with | |- context [measure (algebra.Alg.Term ?f ?t)] => rewrite (measure_equation (algebra.Alg.Term f t)) | H:context [measure (algebra.Alg.Term ?f ?t)] |- _ => rewrite (measure_equation (algebra.Alg.Term f t)) in H|- end ). Definition lt a b := (Zwf.Zwf 0) (marked_measure a) (marked_measure b). Definition le a b := marked_measure a <= marked_measure b. Lemma lt_le_compat : forall a b c, (lt a b) ->(le b c) ->lt a c. Proof. unfold lt, le in *. intros a b c. apply (interp.le_lt_compat_right (interp.o_Z 0)). Qed. Lemma wf_lt : well_founded lt. Proof. unfold lt in *. apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. Qed. Lemma DP_R_xml_0_scc_27_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_27_strict lt. Proof. unfold Relation_Definitions.inclusion, lt in *. intros a b H;destruct H; match goal with | |- (Zwf.Zwf 0) _ (marked_measure (algebra.Alg.Term ?f ?l)) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_lt_compat_right (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma DP_R_xml_0_scc_27_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_27_large le. Proof. unfold Relation_Definitions.inclusion, le, Zwf.Zwf in *. intros a b H;destruct H; match goal with | |- _ <= marked_measure (algebra.Alg.Term ?f ?l) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_trans (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition wf_DP_R_xml_0_scc_27_large := WF_DP_R_xml_0_scc_27_large.wf. Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_27. Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_27_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx';apply DP_R_xml_0_scc_27_strict_in_lt; econstructor eassumption)|| ((apply IHx;[econstructor eassumption| intros y' Hlt;apply IHx';apply lt_le_compat with (1:=Hlt) ; apply DP_R_xml_0_scc_27_large_in_le;econstructor eassumption])). apply wf_DP_R_xml_0_scc_27_large. Qed. End WF_DP_R_xml_0_scc_27. Definition wf_DP_R_xml_0_scc_27 := WF_DP_R_xml_0_scc_27.wf. Lemma acc_DP_R_xml_0_scc_27 : forall x y, (DP_R_xml_0_scc_27 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_27). intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((eapply acc_DP_R_xml_0_non_scc_26; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_25; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). apply wf_DP_R_xml_0_scc_27. Qed. Inductive DP_R_xml_0_non_scc_28 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_28_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_28 (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id_size (x1::nil))::(algebra.Alg.Term algebra.F.id_size (x2::nil))::nil)) (algebra.Alg.Term algebra.F.id_size (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_28 : forall x y, (DP_R_xml_0_non_scc_28 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_27; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_26; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_25; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). Qed. Inductive DP_R_xml_0_non_scc_29 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_29_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_29 (algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id__plus_ ((algebra.Alg.Term algebra.F.id_size (x1::nil))::(algebra.Alg.Term algebra.F.id_size (x2::nil))::nil)):: (algebra.Alg.Term algebra.F.id_1 ((algebra.Alg.Term algebra.F.id__sharp_ nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_size (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_29 : forall x y, (DP_R_xml_0_non_scc_29 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_27; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_26; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_25; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((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_30 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_30_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_scc_30 (algebra.Alg.Term algebra.F.id_size (x2::nil)) (algebra.Alg.Term algebra.F.id_size (x5::nil)) (* *) | DP_R_xml_0_scc_30_1 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_scc_30 (algebra.Alg.Term algebra.F.id_size (x1::nil)) (algebra.Alg.Term algebra.F.id_size (x5::nil)) . Module WF_DP_R_xml_0_scc_30. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_0 (x5:Z) := 2. Definition P_id_bs (x5:Z) := 3 + 3* x5. Definition P_id_and (x5:Z) (x6:Z) := 2* x5. Definition P_id__minus_ (x5:Z) (x6:Z) := 2* x5. Definition P_id_l (x5:Z) := 3 + 1* x5. Definition P_id__plus_ (x5:Z) (x6:Z) := 1* x5 + 1* x6. Definition P_id_wb (x5:Z) := 1. Definition P_id_ge (x5:Z) (x6:Z) := 0. Definition P_id_false := 0. Definition P_id_min (x5:Z) := 3* x5. Definition P_id__sharp_ := 0. Definition P_id_size (x5:Z) := 2 + 2* x5. Definition P_id_if (x5:Z) (x6:Z) (x7:Z) := 1* x6 + 2* x7. Definition P_id_not (x5:Z) := 0. Definition P_id_n (x5:Z) (x6:Z) (x7:Z) := 3 + 3* x5 + 2* x6 + 2* x7. Definition P_id_1 (x5:Z) := 2. Definition P_id_val (x5:Z) := 1* x5. Definition P_id_true := 0. Definition P_id_max (x5:Z) := 1* x5. Lemma P_id_0_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0 x6 <= P_id_0 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_bs x6 <= P_id_bs x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_and x6 x8 <= P_id_and x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__minus_ x6 x8 <= P_id__minus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_l x6 <= P_id_l x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__plus_ x6 x8 <= P_id__plus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_wb x6 <= P_id_wb x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_ge x6 x8 <= P_id_ge x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_min_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_min x6 <= P_id_min x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_size x6 <= P_id_size x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_if x6 x8 x10 <= P_id_if x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_not_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_not x6 <= P_id_not x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_n x6 x8 x10 <= P_id_n x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_1 x6 <= P_id_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_val x6 <= P_id_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_max_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_max x6 <= P_id_max x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : forall x5, (0 <= x5) ->0 <= P_id_0 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_bounded : forall x5, (0 <= x5) ->0 <= P_id_bs x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_and x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__minus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_bounded : forall x5, (0 <= x5) ->0 <= P_id_l x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__plus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_bounded : forall x5, (0 <= x5) ->0 <= P_id_wb x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_ge x6 x5. 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_min_bounded : forall x5, (0 <= x5) ->0 <= P_id_min x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__sharp__bounded : 0 <= P_id__sharp_ . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_bounded : forall x5, (0 <= x5) ->0 <= P_id_size x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_if x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_not_bounded : forall x5, (0 <= x5) ->0 <= P_id_not x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_n x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_1_bounded : forall x5, (0 <= x5) ->0 <= P_id_1 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_bounded : forall x5, (0 <= x5) ->0 <= P_id_val x5. 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_max_bounded : forall x5, (0 <= x5) ->0 <= P_id_max x5. 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_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0 (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_bs (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_and (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) => P_id__minus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_l (x5::nil)) => P_id_l (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) => P_id__plus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_wb (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_ge (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_min (measure x5) | (algebra.Alg.Term algebra.F.id__sharp_ nil) => P_id__sharp_ | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_size (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6::x5::nil)) => P_id_if (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_not (measure x5) | (algebra.Alg.Term algebra.F.id_n (x7::x6::x5::nil)) => P_id_n (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_1 (x5::nil)) => P_id_1 (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_val (measure x5) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_max (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_MIN (x5:Z) := 0. Definition P_id_0_hat_1 (x5:Z) := 0. Definition P_id_SIZE (x5:Z) := 2* x5. Definition P_id_BS (x5:Z) := 0. Definition P_id_GE (x5:Z) (x6:Z) := 0. Definition P_id_IF (x5:Z) (x6:Z) (x7:Z) := 0. Definition P_id_MAX (x5:Z) := 0. Definition P_id__minus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_WB (x5:Z) := 0. Definition P_id__plus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_AND (x5:Z) (x6:Z) := 0. Definition P_id_NOT (x5:Z) := 0. Definition P_id_Marked_val (x5:Z) := 0. Lemma P_id_MIN_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MIN x6 <= P_id_MIN x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_hat_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0_hat_1 x6 <= P_id_0_hat_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SIZE_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_SIZE x6 <= P_id_SIZE x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_BS_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_BS x6 <= P_id_BS x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_GE_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_GE x6 x8 <= P_id_GE x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_IF_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_IF x6 x8 x10 <= P_id_IF x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_MAX_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MAX x6 <= P_id_MAX x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__minus__hat_1 x6 x8 <= P_id__minus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_WB_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_WB x6 <= P_id_WB x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__plus__hat_1 x6 x8 <= P_id__plus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_AND_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_AND x6 x8 <= P_id_AND x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_NOT_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_NOT x6 <= P_id_NOT x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_Marked_val x6 <= P_id_Marked_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max P_id_MIN P_id_0_hat_1 P_id_SIZE P_id_BS P_id_GE P_id_IF P_id_MAX P_id__minus__hat_1 P_id_WB P_id__plus__hat_1 P_id_AND P_id_NOT P_id_Marked_val. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_MIN (measure x5) | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0_hat_1 (measure x5) | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_SIZE (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_BS (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_GE (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6:: x5::nil)) => P_id_IF (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_MAX (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6:: x5::nil)) => P_id__minus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_WB (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6:: x5::nil)) => P_id__plus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_AND (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_NOT (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_Marked_val (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. intros ;apply P_id_MIN_monotonic;assumption. intros ;apply P_id_0_hat_1_monotonic;assumption. intros ;apply P_id_SIZE_monotonic;assumption. intros ;apply P_id_BS_monotonic;assumption. intros ;apply P_id_GE_monotonic;assumption. intros ;apply P_id_IF_monotonic;assumption. intros ;apply P_id_MAX_monotonic;assumption. intros ;apply P_id__minus__hat_1_monotonic;assumption. intros ;apply P_id_WB_monotonic;assumption. intros ;apply P_id__plus__hat_1_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_NOT_monotonic;assumption. intros ;apply P_id_Marked_val_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_30. 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_30. Definition wf_DP_R_xml_0_scc_30 := WF_DP_R_xml_0_scc_30.wf. Lemma acc_DP_R_xml_0_scc_30 : forall x y, (DP_R_xml_0_scc_30 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_30). 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_29; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_28; 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_30. Qed. Inductive DP_R_xml_0_non_scc_31 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_31_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_31 (algebra.Alg.Term algebra.F.id_size (x3::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_31 : forall x y, (DP_R_xml_0_non_scc_31 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_30; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_29; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_28; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). Qed. Inductive DP_R_xml_0_non_scc_32 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_32_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_non_scc_32 (algebra.Alg.Term algebra.F.id_size (x2::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) . Lemma acc_DP_R_xml_0_non_scc_32 : forall x y, (DP_R_xml_0_non_scc_32 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_30; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_29; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_28; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). Qed. Inductive DP_R_xml_0_scc_33 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_33_0 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_scc_33 (algebra.Alg.Term algebra.F.id_wb (x3::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) (* *) | DP_R_xml_0_scc_33_1 : forall x2 x1 x5 x3, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_n (x1::x2::x3::nil)) x5) -> DP_R_xml_0_scc_33 (algebra.Alg.Term algebra.F.id_wb (x2::nil)) (algebra.Alg.Term algebra.F.id_wb (x5::nil)) . Module WF_DP_R_xml_0_scc_33. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_0 (x5:Z) := 0. Definition P_id_bs (x5:Z) := 3* x5. Definition P_id_and (x5:Z) (x6:Z) := 2* x5. Definition P_id__minus_ (x5:Z) (x6:Z) := 1* x5 + 2* x6. Definition P_id_l (x5:Z) := 2 + 1* x5. Definition P_id__plus_ (x5:Z) (x6:Z) := 1* x5 + 1* x6. Definition P_id_wb (x5:Z) := 2 + 2* x5. Definition P_id_ge (x5:Z) (x6:Z) := 1. Definition P_id_false := 0. Definition P_id_min (x5:Z) := 3* x5. Definition P_id__sharp_ := 0. Definition P_id_size (x5:Z) := 1* x5. Definition P_id_if (x5:Z) (x6:Z) (x7:Z) := 1* x5 + 2* x6 + 1* x7. Definition P_id_not (x5:Z) := 1. Definition P_id_n (x5:Z) (x6:Z) (x7:Z) := 3 + 3* x5 + 3* x6 + 2* x7. Definition P_id_1 (x5:Z) := 2. Definition P_id_val (x5:Z) := 1* x5. Definition P_id_true := 1. Definition P_id_max (x5:Z) := 1* x5. Lemma P_id_0_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0 x6 <= P_id_0 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_bs x6 <= P_id_bs x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_and x6 x8 <= P_id_and x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__minus_ x6 x8 <= P_id__minus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_l x6 <= P_id_l x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id__plus_ x6 x8 <= P_id__plus_ x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_wb x6 <= P_id_wb x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_ge x6 x8 <= P_id_ge x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_min_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_min x6 <= P_id_min x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_size x6 <= P_id_size x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_if x6 x8 x10 <= P_id_if x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_not_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_not x6 <= P_id_not x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_n x6 x8 x10 <= P_id_n x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_1 x6 <= P_id_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_val x6 <= P_id_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_max_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_max x6 <= P_id_max x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : forall x5, (0 <= x5) ->0 <= P_id_0 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_bs_bounded : forall x5, (0 <= x5) ->0 <= P_id_bs x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_and_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_and x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__minus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_l_bounded : forall x5, (0 <= x5) ->0 <= P_id_l x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id__plus_ x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_wb_bounded : forall x5, (0 <= x5) ->0 <= P_id_wb x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ge_bounded : forall x6 x5, (0 <= x5) ->(0 <= x6) ->0 <= P_id_ge x6 x5. 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_min_bounded : forall x5, (0 <= x5) ->0 <= P_id_min x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__sharp__bounded : 0 <= P_id__sharp_ . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_size_bounded : forall x5, (0 <= x5) ->0 <= P_id_size x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_if x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_not_bounded : forall x5, (0 <= x5) ->0 <= P_id_not x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_n_bounded : forall x6 x5 x7, (0 <= x5) ->(0 <= x6) ->(0 <= x7) ->0 <= P_id_n x7 x6 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_1_bounded : forall x5, (0 <= x5) ->0 <= P_id_1 x5. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_val_bounded : forall x5, (0 <= x5) ->0 <= P_id_val x5. 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_max_bounded : forall x5, (0 <= x5) ->0 <= P_id_max x5. 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_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0 (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_bs (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_and (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6::x5::nil)) => P_id__minus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_l (x5::nil)) => P_id_l (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6::x5::nil)) => P_id__plus_ (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_wb (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_ge (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_min (measure x5) | (algebra.Alg.Term algebra.F.id__sharp_ nil) => P_id__sharp_ | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_size (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6::x5::nil)) => P_id_if (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_not (measure x5) | (algebra.Alg.Term algebra.F.id_n (x7::x6::x5::nil)) => P_id_n (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_1 (x5::nil)) => P_id_1 (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_val (measure x5) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_max (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_MIN (x5:Z) := 0. Definition P_id_0_hat_1 (x5:Z) := 0. Definition P_id_SIZE (x5:Z) := 0. Definition P_id_BS (x5:Z) := 0. Definition P_id_GE (x5:Z) (x6:Z) := 0. Definition P_id_IF (x5:Z) (x6:Z) (x7:Z) := 0. Definition P_id_MAX (x5:Z) := 0. Definition P_id__minus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_WB (x5:Z) := 2* x5. Definition P_id__plus__hat_1 (x5:Z) (x6:Z) := 0. Definition P_id_AND (x5:Z) (x6:Z) := 0. Definition P_id_NOT (x5:Z) := 0. Definition P_id_Marked_val (x5:Z) := 0. Lemma P_id_MIN_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MIN x6 <= P_id_MIN x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_hat_1_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_0_hat_1 x6 <= P_id_0_hat_1 x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SIZE_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_SIZE x6 <= P_id_SIZE x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_BS_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_BS x6 <= P_id_BS x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_GE_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_GE x6 x8 <= P_id_GE x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_IF_monotonic : forall x8 x10 x6 x9 x5 x7, (0 <= x10)/\ (x10 <= x9) -> (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_IF x6 x8 x10 <= P_id_IF x5 x7 x9. Proof. intros x10 x9 x8 x7 x6 x5. 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_MAX_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_MAX x6 <= P_id_MAX x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__minus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__minus__hat_1 x6 x8 <= P_id__minus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_WB_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_WB x6 <= P_id_WB x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id__plus__hat_1_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) -> P_id__plus__hat_1 x6 x8 <= P_id__plus__hat_1 x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_AND_monotonic : forall x8 x6 x5 x7, (0 <= x8)/\ (x8 <= x7) -> (0 <= x6)/\ (x6 <= x5) ->P_id_AND x6 x8 <= P_id_AND x5 x7. Proof. intros x8 x7 x6 x5. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_NOT_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_NOT x6 <= P_id_NOT x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_val_monotonic : forall x6 x5, (0 <= x6)/\ (x6 <= x5) ->P_id_Marked_val x6 <= P_id_Marked_val x5. Proof. intros x6 x5. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_0 P_id_bs P_id_and P_id__minus_ P_id_l P_id__plus_ P_id_wb P_id_ge P_id_false P_id_min P_id__sharp_ P_id_size P_id_if P_id_not P_id_n P_id_1 P_id_val P_id_true P_id_max P_id_MIN P_id_0_hat_1 P_id_SIZE P_id_BS P_id_GE P_id_IF P_id_MAX P_id__minus__hat_1 P_id_WB P_id__plus__hat_1 P_id_AND P_id_NOT P_id_Marked_val. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_min (x5::nil)) => P_id_MIN (measure x5) | (algebra.Alg.Term algebra.F.id_0 (x5::nil)) => P_id_0_hat_1 (measure x5) | (algebra.Alg.Term algebra.F.id_size (x5::nil)) => P_id_SIZE (measure x5) | (algebra.Alg.Term algebra.F.id_bs (x5::nil)) => P_id_BS (measure x5) | (algebra.Alg.Term algebra.F.id_ge (x6::x5::nil)) => P_id_GE (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_if (x7::x6:: x5::nil)) => P_id_IF (measure x7) (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_max (x5::nil)) => P_id_MAX (measure x5) | (algebra.Alg.Term algebra.F.id__minus_ (x6:: x5::nil)) => P_id__minus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_wb (x5::nil)) => P_id_WB (measure x5) | (algebra.Alg.Term algebra.F.id__plus_ (x6:: x5::nil)) => P_id__plus__hat_1 (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_and (x6::x5::nil)) => P_id_AND (measure x6) (measure x5) | (algebra.Alg.Term algebra.F.id_not (x5::nil)) => P_id_NOT (measure x5) | (algebra.Alg.Term algebra.F.id_val (x5::nil)) => P_id_Marked_val (measure x5) | _ => 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_0_monotonic;assumption. intros ;apply P_id_bs_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id__minus__monotonic;assumption. intros ;apply P_id_l_monotonic;assumption. intros ;apply P_id__plus__monotonic;assumption. intros ;apply P_id_wb_monotonic;assumption. intros ;apply P_id_ge_monotonic;assumption. intros ;apply P_id_min_monotonic;assumption. intros ;apply P_id_size_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_n_monotonic;assumption. intros ;apply P_id_1_monotonic;assumption. intros ;apply P_id_val_monotonic;assumption. intros ;apply P_id_max_monotonic;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_bs_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id__minus__bounded;assumption. intros ;apply P_id_l_bounded;assumption. intros ;apply P_id__plus__bounded;assumption. intros ;apply P_id_wb_bounded;assumption. intros ;apply P_id_ge_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_min_bounded;assumption. intros ;apply P_id__sharp__bounded;assumption. intros ;apply P_id_size_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_n_bounded;assumption. intros ;apply P_id_1_bounded;assumption. intros ;apply P_id_val_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_max_bounded;assumption. apply rules_monotonic. intros ;apply P_id_MIN_monotonic;assumption. intros ;apply P_id_0_hat_1_monotonic;assumption. intros ;apply P_id_SIZE_monotonic;assumption. intros ;apply P_id_BS_monotonic;assumption. intros ;apply P_id_GE_monotonic;assumption. intros ;apply P_id_IF_monotonic;assumption. intros ;apply P_id_MAX_monotonic;assumption. intros ;apply P_id__minus__hat_1_monotonic;assumption. intros ;apply P_id_WB_monotonic;assumption. intros ;apply P_id__plus__hat_1_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_NOT_monotonic;assumption. intros ;apply P_id_Marked_val_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_33. Proof. intros x. apply well_founded_ind with (R:=fun x y => (Zwf.Zwf 0) (marked_measure x) (marked_measure y)). apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. clear x. intros x IHx. repeat ( constructor;inversion 1;subst; full_prove_ineq algebra.Alg.Term ltac:(algebra.Alg_ext.find_replacement ) algebra.EQT_ext.one_step_list_refl_trans_clos marked_measure marked_measure_star_monotonic (Zwf.Zwf 0) (interp.o_Z 0) ltac:(fun _ => R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 ) ltac:(fun _ => rewrite_and_unfold ) ltac:(fun _ => generate_pos_hyp ) ltac:(fun _ => cbv beta iota zeta delta - [Zplus Zmult Zle Zlt] in * ; try (constructor)) IHx ). Qed. End WF_DP_R_xml_0_scc_33. Definition wf_DP_R_xml_0_scc_33 := WF_DP_R_xml_0_scc_33.wf. Lemma acc_DP_R_xml_0_scc_33 : forall x y, (DP_R_xml_0_scc_33 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_33). intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((eapply acc_DP_R_xml_0_non_scc_32; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_31; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_24; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_23; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_16; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_15; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_14; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_2; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))))))))))). apply wf_DP_R_xml_0_scc_33. 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_32; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_31; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_30; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_29; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_28; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_27; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_26; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_25; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_24; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_23; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_22; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_21; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_20; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_19; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_18; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_17; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_16; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_15; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_14; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_13; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_12; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_11; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_10; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_9; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_8; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_7; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_6; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_5; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_4; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_3; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_2; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_1; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_0; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_33; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_32; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_31; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_30; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_29; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_28; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_27; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_26; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_25; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_24; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_23; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_22; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_21; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_20; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_19; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_18; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_17; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_16; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_15; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_14; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_13; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_12; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_11; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_10; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_9; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_8; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_7; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_6; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_5; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_4; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| (( eapply acc_DP_R_xml_0_scc_3; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_2; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_1; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_0; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| ( fail)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))). Qed. End WF_DP_R_xml_0. Definition wf_H := WF_DP_R_xml_0.wf. Lemma wf : well_founded (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules). Proof. apply ddp.dp_criterion. apply R_xml_0_deep_rew.R_xml_0_non_var. apply R_xml_0_deep_rew.R_xml_0_reg. intros ; apply (ddp.constructor_defined_dec _ _ R_xml_0_deep_rew.R_xml_0_rules_included). refine (Inclusion.wf_incl _ _ _ _ wf_H). intros x y H. destruct (R_xml_0_dp_step_spec H) as [f [l1 [l2 [H1 [H2 H3]]]]]. destruct (ddp.dp_list_complete _ _ R_xml_0_deep_rew.R_xml_0_rules_included _ _ H3) as [x' [y' [sigma [h1 [h2 h3]]]]]. clear H3. subst. vm_compute in h3|-. let e := type of h3 in (dp_concl_tac h2 h3 ltac:(fun _ => idtac) e). Qed. End WF_R_xml_0_deep_rew. (* *** Local Variables: *** *** coq-prog-name: "coqtop" *** *** coq-prog-args: ("-emacs-U" "-I" "$COCCINELLE/examples" "-I" "$COCCINELLE/term_algebra" "-I" "$COCCINELLE/term_orderings" "-I" "$COCCINELLE/basis" "-I" "$COCCINELLE/list_extensions" "-I" "$COCCINELLE/examples/cime_trace/") *** *** compile-command: "coqc -I $COCCINELLE/term_algebra -I $COCCINELLE/term_orderings -I $COCCINELLE/basis -I $COCCINELLE/list_extensions -I $COCCINELLE/examples/cime_trace/ -I $COCCINELLE/examples/ c_output/strat/tpdb-5.0___TRS___Cime___tree.trs/a3pat.v" *** *** End: *** *)